diff --git a/thys/DynamicArchitectures/Configuration_Traces.thy b/thys/DynamicArchitectures/Configuration_Traces.thy --- a/thys/DynamicArchitectures/Configuration_Traces.thy +++ b/thys/DynamicArchitectures/Configuration_Traces.thy @@ -1,1803 +1,1804 @@ (* Title: Configuration_Traces.thy Author: Diego Marmsoler *) section "A Theory of Dynamic Architectures" text \ The following theory formalizes configuration traces~\cite{Marmsoler2016a,Marmsoler2016} as a model for dynamic architectures. Since configuration traces may be finite as well as infinite, the theory depends on Lochbihler's theory of co-inductive lists~\cite{Lochbihler2010}. \ theory Configuration_Traces imports Coinductive.Coinductive_List begin text \ In the following we first provide some preliminary results for natural numbers, extended natural numbers, and lazy lists. Then, we introduce a locale @text{dynamic\_architectures} which introduces basic definitions and corresponding properties for dynamic architectures. \ subsection "Natural Numbers" text \ We provide one additional property for natural numbers. \ lemma boundedGreatest: assumes "P (i::nat)" and "\n' > n. \ P n'" shows "\i'\n. P i' \ (\n'. P n' \ n'\i')" proof - have "P (i::nat) \ n\i \ \n' > n. \ P n' \ (\i'\n. P i' \ (\n'\n. P n' \ n'\i'))" proof (induction n) case 0 then show ?case by auto next case (Suc n) then show ?case proof cases assume "i = Suc n" then show ?thesis using Suc.prems by auto next assume "\(i = Suc n)" thus ?thesis proof cases assume "P (Suc n)" thus ?thesis by auto next assume "\ P (Suc n)" with Suc.prems have "\n' > n. \ P n'" using Suc_lessI by blast moreover from \\(i = Suc n)\ have "i \ n" and "P i" using Suc.prems by auto ultimately obtain i' where "i'\n \ P i' \ (\n'\n. P n' \ n' \ i')" using Suc.IH by blast hence "i' \ n" and "P i'" and "(\n'\n. P n' \ n' \ i')" by auto thus ?thesis by (metis le_SucI le_Suc_eq) qed qed qed moreover have "n\i" proof (rule ccontr) assume "\ (n \ i)" hence "n < i" by arith thus False using assms by blast qed ultimately obtain i' where "i'\n" and "P i'" and "\n'\n. P n' \ n' \ i'" using assms by blast with assms have "\n'. P n' \ n' \ i'" using not_le_imp_less by blast with \i' \ n\ and \P i'\ show ?thesis by auto qed subsection "Extended Natural Numbers" text \ We provide one simple property for the \emph{strict} order over extended natural numbers. \ lemma enat_min: assumes "m \ enat n'" and "enat n < m - enat n'" shows "enat n + enat n' < m" using assms by (metis add.commute enat.simps(3) enat_add_mono enat_add_sub_same le_iff_add) subsection "Lazy Lists" text \ In the following we provide some additional notation and properties for lazy lists. \ notation LNil ("[]\<^sub>l") notation LCons (infixl "#\<^sub>l" 60) notation lappend (infixl "@\<^sub>l" 60) lemma lnth_lappend[simp]: assumes "lfinite xs" and "\ lnull ys" shows "lnth (xs @\<^sub>l ys) (the_enat (llength xs)) = lhd ys" proof - from assms have "\k. llength xs = enat k" using lfinite_conv_llength_enat by auto then obtain k where "llength xs = enat k" by blast hence "lnth (xs @\<^sub>l ys) (the_enat (llength xs)) = lnth ys 0" using lnth_lappend2[of xs k k ys] by simp with assms show ?thesis using lnth_0_conv_lhd by simp qed lemma lfilter_ltake: assumes "\(n::nat)\llength xs. n\i \ (\ P (lnth xs n))" shows "lfilter P xs = lfilter P (ltake i xs)" proof - have "lfilter P xs = lfilter P ((ltake i xs) @\<^sub>l (ldrop i xs))" using lappend_ltake_ldrop[of "(enat i)" xs] by simp hence "lfilter P xs = (lfilter P ((ltake i) xs)) @\<^sub>l (lfilter P (ldrop i xs))" by simp show ?thesis proof cases assume "enat i \ llength xs" have "\x P (lnth (ldrop i xs) x)" proof (rule allI) fix x show "enat x < llength (ldrop (enat i) xs) \ \ P (lnth (ldrop (enat i) xs) x)" proof assume "enat x < llength (ldrop (enat i) xs)" moreover have "llength (ldrop (enat i) xs) = llength xs - enat i" using llength_ldrop[of "enat i"] by simp ultimately have "enat x < llength xs - enat i" by simp with \enat i \ llength xs\ have "enat x + enat i < llength xs" using enat_min[of i "llength xs" x] by simp moreover have "enat i + enat x = enat x + enat i" by simp ultimately have "enat i + enat x < llength xs" by arith hence "i + x < llength xs" by simp hence "lnth (ldrop i xs) x = lnth xs (x + the_enat i)" using lnth_ldrop[of "enat i" "x" xs] by simp moreover have "x + i \ i" by simp with assms \i + x < llength xs\ have "\ P (lnth xs (x + the_enat i))" by (simp add: assms(1) add.commute) ultimately show "\ P (lnth (ldrop i xs) x)" using assms by simp qed qed hence "lfilter P (ldrop i xs) = []\<^sub>l" by (metis diverge_lfilter_LNil in_lset_conv_lnth) with \lfilter P xs = (lfilter P ((ltake i) xs)) @\<^sub>l (lfilter P (ldrop i xs))\ show "lfilter P xs = lfilter P (ltake i xs)" by simp next assume "\ enat i \ llength xs" hence "enat i > llength xs" by simp hence "ldrop i xs = []\<^sub>l" by simp hence "lfilter P (ldrop i xs) = []\<^sub>l" using lfilter_LNil[of P] by arith with \lfilter P xs = (lfilter P ((ltake i) xs)) @\<^sub>l (lfilter P (ldrop i xs))\ show "lfilter P xs = lfilter P (ltake i xs)" by simp qed qed lemma lfilter_lfinite[simp]: assumes "lfinite (lfilter P t)" and "\ lfinite t" shows "\n. \n'>n. \ P (lnth t n')" proof - from assms have "finite {n. enat n < llength t \ P (lnth t n)}" using lfinite_lfilter by auto then obtain k where sset: "{n. enat n < llength t \ P (lnth t n)} \ {n. n enat n < llength t \ P (lnth t n)}" using finite_nat_bounded[of "{n. enat n < llength t \ P (lnth t n)}"] by auto show ?thesis proof (rule ccontr) assume "\(\n. \n'>n. \ P (lnth t n'))" hence "\n. \n'>n. P (lnth t n')" by simp then obtain n' where "n'>k" and "P (lnth t n')" by auto moreover from \\ lfinite t\ have "n' < llength t" by (simp add: not_lfinite_llength) ultimately have "n' \ {n. n enat n < llength t \ P (lnth t n)}" and "n'\{n. enat n < llength t \ P (lnth t n)}" by auto with sset show False by auto qed qed subsection "Specifying Dynamic Architectures" text \ In the following we formalize dynamic architectures in terms of configuration traces, i.e., sequences of architecture configurations. Moreover, we introduce definitions for operations to support the specification of configuration traces. \ typedecl cnf type_synonym trace = "nat \ cnf" consts arch:: "trace set" type_synonym cta = "trace \ nat \ bool" subsubsection "Implication" definition imp :: "cta \ cta \ cta" (infixl "\\<^sup>c" 10) where "\ \\<^sup>c \' \ \ t n. \ t n \ \' t n" declare imp_def[simp] lemma impI[intro!]: fixes t n assumes "\ t n \ \' t n" shows "(\ \\<^sup>c \') t n" using assms by simp lemma impE[elim!]: fixes t n assumes "(\ \\<^sup>c \') t n" and "\ t n" and "\' t n \ \'' t n" shows "\'' t n" using assms by simp subsubsection "Disjunction" definition disj :: "cta \ cta \ cta" (infixl "\\<^sup>c" 15) where "\ \\<^sup>c \' \ \ t n. \ t n \ \' t n" declare disj_def[simp] lemma disjI1[intro]: assumes "\ t n" shows "(\ \\<^sup>c \') t n" using assms by simp lemma disjI2[intro]: assumes "\' t n" shows "(\ \\<^sup>c \') t n" using assms by simp lemma disjE[elim!]: assumes "(\ \\<^sup>c \') t n" and "\ t n \ \'' t n" and "\' t n \ \'' t n" shows "\'' t n" using assms by auto subsubsection "Conjunction" definition conj :: "cta \ cta \ cta" (infixl "\\<^sup>c" 20) where "\ \\<^sup>c \' \ \ t n. \ t n \ \' t n" declare conj_def[simp] lemma conjI[intro!]: fixes n assumes "\ t n" and "\' t n" shows "(\ \\<^sup>c \') t n" using assms by simp lemma conjE[elim!]: fixes n assumes "(\ \\<^sup>c \') t n" and "\ t n \ \' t n \ \'' t n" shows "\'' t n" using assms by simp subsubsection "Negation" definition neg :: "cta \ cta" ("\\<^sup>c _" [19] 19) where "\\<^sup>c \ \ \ t n. \ \ t n" declare neg_def[simp] lemma negI[intro!]: assumes "\ t n \ False" shows "(\\<^sup>c \) t n" using assms by auto lemma negE[elim!]: assumes "(\\<^sup>c \) t n" and "\ t n" shows "\' t n" using assms by simp subsubsection "Quantifiers" definition all :: "('a \ cta) \ cta" (binder "\\<^sub>c" 10) where "all P \ \t n. (\y. (P y t n))" declare all_def[simp] lemma allI[intro!]: assumes "\x. \ x t n" shows "(\\<^sub>cx. \ x) t n" using assms by simp lemma allE[elim!]: fixes n assumes "(\\<^sub>cx. \ x) t n" and "\ x t n \ \' t n" shows "\' t n" using assms by simp definition ex :: "('a \ cta) \ cta" (binder "\\<^sub>c" 10) where "ex P \ \t n. (\y. (P y t n))" declare ex_def[simp] lemma exI[intro!]: assumes "\ x t n" shows "(\\<^sub>cx. \ x) t n" using assms HOL.exI by simp lemma exE[elim!]: assumes "(\\<^sub>cx. \ x) t n" and "\x. \ x t n \ \' t n" shows "\' t n" using assms HOL.exE by auto subsubsection "Atomic Assertions" text \ First we provide rules for basic behavior assertions. \ definition ca :: "(cnf \ bool) \ cta" where "ca \ \ \ t n. \ (t n)" lemma caI[intro]: fixes n assumes "\ (t n)" shows "(ca \) t n" using assms ca_def by simp lemma caE[elim]: fixes n assumes "(ca \) t n" shows "\ (t n)" using assms ca_def by simp subsubsection "Next Operator" definition nxt :: "cta \ cta" ("\\<^sub>c(_)" 24) where "\\<^sub>c(\) \ \(t::(nat \ cnf)) n. \ t (Suc n)" subsubsection "Eventually Operator" definition evt :: "cta \ cta" ("\\<^sub>c(_)" 23) where "\\<^sub>c(\) \ \(t::(nat \ cnf)) n. \n'\n. \ t n'" subsubsection "Globally Operator" definition glob :: "cta \ cta" ("\\<^sub>c(_)" 22) where "\\<^sub>c(\) \ \(t::(nat \ cnf)) n. \n'\n. \ t n'" lemma globI[intro!]: fixes n' assumes "\n\n'. \ t n" shows "(\\<^sub>c(\)) t n'" using assms glob_def by simp lemma globE[elim!]: fixes n n' assumes "(\\<^sub>c(\)) t n" and "n'\n" shows "\ t n'" using assms glob_def by simp subsubsection "Until Operator" definition until :: "cta \ cta \ cta" (infixl "\\<^sub>c" 21) where "\' \\<^sub>c \ \ \(t::(nat \ cnf)) n. \n''\n. \ t n'' \ (\n'\n. n' < n'' \ \' t n')" lemma untilI[intro]: fixes n assumes "\n''\n. \ t n'' \ (\n'\n. n' \' t n')" shows "(\' \\<^sub>c \) t n" using assms until_def by simp lemma untilE[elim]: fixes n assumes "(\' \\<^sub>c \) t n" shows "\n''\n. \ t n'' \ (\n'\n. n' \' t n')" using assms until_def by simp subsubsection "Weak Until" definition wuntil :: "cta \ cta \ cta" (infixl "\\<^sub>c" 20) where "\' \\<^sub>c \ \ \' \\<^sub>c \ \\<^sup>c \\<^sub>c(\')" lemma wUntilI[intro]: fixes n assumes "(\n''\n. \ t n'' \ (\n'\n. n' \' t n')) \ (\n'\n. \' t n')" shows "(\' \\<^sub>c \) t n" using assms wuntil_def by auto lemma wUntilE[elim]: fixes n n' assumes "(\' \\<^sub>c \) t n" shows "(\n''\n. \ t n'' \ (\n'\n. n' \' t n')) \ (\n'\n. \' t n')" proof - from assms have "(\' \\<^sub>c \ \\<^sup>c \\<^sub>c(\')) t n" using wuntil_def by simp hence "(\' \\<^sub>c \) t n \ (\\<^sub>c(\')) t n" by simp thus ?thesis proof assume "(\' \\<^sub>c \) t n" hence "\n''\n. \ t n'' \ (\n'\n. n' < n'' \ \' t n')" by auto thus ?thesis by auto next assume "(\\<^sub>c\') t n" hence "\n'\n. \' t n'" by auto thus ?thesis by auto qed qed lemma wUntil_Glob: assumes "(\' \\<^sub>c \) t n" and "(\\<^sub>c(\' \\<^sup>c \'')) t n" shows "(\'' \\<^sub>c \) t n" proof from assms(1) have "(\n''\n. \ t n'' \ (\n'\n. n' < n'' \ \' t n')) \ (\n'\n. \' t n')" using wUntilE by simp thus "(\n''\n. \ t n'' \ (\n'\n. n' < n'' \ \'' t n')) \ (\n'\n. \'' t n')" proof assume "\n''\n. \ t n'' \ (\n'\n. n' < n'' \ \' t n')" show "(\n''\n. \ t n'' \ (\n'\n. n' < n'' \ \'' t n')) \ (\n'\n. \'' t n')" proof - from \\n''\n. \ t n'' \ (\n'\n. n' < n'' \ \' t n')\ obtain n'' where "n''\n" and "\ t n''" and a1: "\n'\n. n' < n'' \ \' t n'" by auto moreover have "\n'\n. n' < n'' \ \'' t n'" proof fix n' show "n'\n \ n'< n'' \ \'' t n'" proof (rule HOL.impI[OF HOL.impI]) assume "n'\n" and "n'' \\<^sup>c \'') t n'" using globE by simp hence "\' t n' \ \'' t n'" using impE by auto moreover from a1 \n'\n\ \n' have "\' t n'" by simp ultimately show "\'' t n'" by simp qed qed ultimately show ?thesis by auto qed next assume a1: "\n'\n. \' t n'" have "\n'\n. \'' t n'" proof fix n' show "n'\n \ \'' t n'" proof assume "n'\n" with assms(2) have "(\' \\<^sup>c \'') t n'" using globE by simp hence "\' t n' \ \'' t n'" using impE by auto moreover from a1 \n'\n\ have "\' t n'" by simp ultimately show "\'' t n'" by simp qed qed thus "(\n''\n. \ t n'' \ (\n'\n. n' < n'' \ \'' t n')) \ (\n'\n. \'' t n')" by simp qed qed subsection "Dynamic Components" text \ To support the specification of patterns over dynamic architectures we provide a locale for dynamic components. It takes the following type parameters: \begin{itemize} \item id: a type for component identifiers \item cmp: a type for components \item cnf: a type for architecture configurations \end{itemize} \ locale dynamic_component = fixes tCMP :: "'id \ cnf \ 'cmp" ("\\<^bsub>_\<^esub>(_)" [0,110]60) and active :: "'id \ cnf \ bool" ("\_\\<^bsub>_\<^esub>" [0,110]60) begin text \ The locale requires two parameters: \begin{itemize} \item @{term tCMP} is an operator to obtain a component with a certain identifier from an architecture configuration. \item @{term active} is a predicate to assert whether a certain component is activated within an architecture configuration. \end{itemize} \ text \ The locale provides some general properties about its parameters and introduces six important operators over configuration traces: \begin{itemize} \item An operator to extract the behavior of a certain component out of a given configuration trace. \item An operator to obtain the number of activations of a certain component within a given configuration trace. \item An operator to obtain the least point in time (before a certain point in time) from which on a certain component is not activated anymore. \item An operator to obtain the latest point in time where a certain component was activated. \item Two operators to map time-points between configuration traces and behavior traces. \end{itemize} Moreover, the locale provides several properties about the operators and their relationships. \ lemma nact_active: fixes t::"nat \ cnf" and n::nat and n'' and id assumes "\id\\<^bsub>t n\<^esub>" and "n'' \ n" and "\ (\n'\n. n' < n'' \ \id\\<^bsub>t n'\<^esub>)" shows "n=n''" using assms le_eq_less_or_eq by auto lemma nact_exists: fixes t::"nat \ cnf" assumes "\i\n. \c\\<^bsub>t i\<^esub>" shows "\i\n. \c\\<^bsub>t i\<^esub> \ (\k. n\k \ k \c\\<^bsub>t k\<^esub>)" proof - let ?L = "LEAST i. (i\n \ \c\\<^bsub>t i\<^esub>)" from assms have "?L\n \ \c\\<^bsub>t ?L\<^esub>" using LeastI[of "\x::nat. (x\n \ \c\\<^bsub>t x\<^esub>)"] by auto moreover have "\k. n\k \ k \c\\<^bsub>t k\<^esub>" using not_less_Least by auto ultimately show ?thesis by blast qed lemma lActive_least: assumes "\i\n. i < llength t \ \c\\<^bsub>lnth t i\<^esub>" shows "\i\n. (i < llength t \ \c\\<^bsub>lnth t i\<^esub> \ (\k. n\k \ k k \c\\<^bsub>lnth t k\<^esub>))" proof - let ?L = "LEAST i. (i\n \ i < llength t \ \c\\<^bsub>lnth t i\<^esub>)" from assms have "?L\n \ ?L < llength t \ \c\\<^bsub>lnth t ?L\<^esub>" using LeastI[of "\x::nat.(x\n \ x \c\\<^bsub>lnth t x\<^esub>)"] by auto moreover have "\k. n\k \ k k \c\\<^bsub>lnth t k\<^esub>" using not_less_Least by auto ultimately show ?thesis by blast qed subsection "Projection" text \ In the following we introduce an operator which extracts the behavior of a certain component out of a given configuration trace. \ definition proj:: "'id \ (cnf llist) \ ('cmp llist)" ("\\<^bsub>_\<^esub>(_)" [0,110]60) where "proj c = lmap (\cnf. (\\<^bsub>c\<^esub>(cnf))) \ (lfilter (active c))" lemma proj_lnil [simp,intro]: "\\<^bsub>c\<^esub>([]\<^sub>l) = []\<^sub>l" using proj_def by simp lemma proj_lnull [simp]: "\\<^bsub>c\<^esub>(t) = []\<^sub>l \ (\k \ lset t. \ \c\\<^bsub>k\<^esub>)" proof assume "\\<^bsub>c\<^esub>(t) = []\<^sub>l" hence "lfilter (active c) t = []\<^sub>l" using proj_def lmap_eq_LNil by auto thus "\k \ lset t. \ \c\\<^bsub>k\<^esub>" using lfilter_eq_LNil[of "active c"] by simp next assume "\k\lset t. \ \c\\<^bsub>k\<^esub>" hence "lfilter (active c) t = []\<^sub>l" by simp thus "\\<^bsub>c\<^esub>(t) = []\<^sub>l" using proj_def by simp qed lemma proj_LCons [simp]: "\\<^bsub>i\<^esub>(x #\<^sub>l xs) = (if \i\\<^bsub>x\<^esub> then (\\<^bsub>i\<^esub>(x)) #\<^sub>l (\\<^bsub>i\<^esub>(xs)) else \\<^bsub>i\<^esub>(xs))" using proj_def by simp lemma proj_llength[simp]: "llength (\\<^bsub>c\<^esub>(t)) \ llength t" using llength_lfilter_ile proj_def by simp lemma proj_ltake: assumes "\(n'::nat)\llength t. n'\n \ (\ \c\\<^bsub>lnth t n'\<^esub>)" shows "\\<^bsub>c\<^esub>(t) = \\<^bsub>c\<^esub>(ltake n t)" using lfilter_ltake proj_def assms by (metis comp_apply) lemma proj_finite_bound: assumes "lfinite (\\<^bsub>c\<^esub>(inf_llist t))" shows "\n. \n'>n. \ \c\\<^bsub>t n'\<^esub>" using assms lfilter_lfinite[of "active c" "inf_llist t"] proj_def by simp subsubsection "Monotonicity and Continuity" lemma proj_mcont: shows "mcont lSup lprefix lSup lprefix (proj c)" proof - have "mcont lSup lprefix lSup lprefix (\x. lmap (\cnf. \\<^bsub>c\<^esub>(cnf)) (lfilter (active c) x))" by simp moreover have "(\x. lmap (\cnf. \\<^bsub>c\<^esub>(cnf)) (lfilter (active c) x)) = lmap (\cnf. \\<^bsub>c\<^esub>(cnf)) \ lfilter (active c)" by auto ultimately show ?thesis using proj_def by simp qed lemma proj_mcont2mcont: assumes "mcont lub ord lSup lprefix f" shows "mcont lub ord lSup lprefix (\x. \\<^bsub>c\<^esub>(f x))" proof - have "mcont lSup lprefix lSup lprefix (proj c)" using proj_mcont by simp with assms show ?thesis using llist.mcont2mcont[of lSup lprefix "proj c"] by simp qed lemma proj_mono_prefix[simp]: assumes "lprefix t t'" shows "lprefix (\\<^bsub>c\<^esub>(t)) (\\<^bsub>c\<^esub>(t'))" proof - from assms have "lprefix (lfilter (active c) t) (lfilter (active c) t')" using lprefix_lfilterI by simp hence "lprefix (lmap (\cnf. \\<^bsub>c\<^esub>(cnf)) (lfilter (active c) t)) (lmap (\cnf. \\<^bsub>c\<^esub>(cnf)) (lfilter (active c) t'))" using lmap_lprefix by simp thus ?thesis using proj_def by simp qed subsubsection "Finiteness" lemma proj_finite[simp]: assumes "lfinite t" shows "lfinite (\\<^bsub>c\<^esub>(t))" using assms proj_def by simp lemma proj_finite2: assumes "\(n'::nat)\llength t. n'\n \ (\ \c\\<^bsub>lnth t n'\<^esub>)" shows "lfinite (\\<^bsub>c\<^esub>(t))" using assms proj_ltake proj_finite by simp lemma proj_append_lfinite[simp]: fixes t t' assumes "lfinite t" shows "\\<^bsub>c\<^esub>(t @\<^sub>l t') = (\\<^bsub>c\<^esub>(t)) @\<^sub>l (\\<^bsub>c\<^esub>(t'))" (is "?lhs=?rhs") proof - have "?lhs = (lmap (\cnf. \\<^bsub>c\<^esub>(cnf)) \ (lfilter (active c))) (t @\<^sub>l t')" using proj_def by simp also have "\ = lmap (\cnf. \\<^bsub>c\<^esub>(cnf)) (lfilter (active c) (t @\<^sub>l t'))" by simp also from assms have "\ = lmap (\cnf. \\<^bsub>c\<^esub>(cnf)) ((lfilter (active c) t) @\<^sub>l (lfilter (active c) t'))" by simp also have "\ = (@\<^sub>l) (lmap (\cnf. \\<^bsub>c\<^esub>(cnf)) (lfilter (active c) t)) (lmap (\cnf. \\<^bsub>c\<^esub>(cnf)) (lfilter (active c) t'))" using lmap_lappend_distrib by simp also have "\ = ?rhs" using proj_def by simp finally show ?thesis . qed lemma proj_one: assumes "\i. i < llength t \ \c\\<^bsub>lnth t i\<^esub>" shows "llength (\\<^bsub>c\<^esub>(t)) \ 1" proof - from assms have "\x\lset t. \c\\<^bsub>x\<^esub>" using lset_conv_lnth by force hence "\ lfilter (\k. \c\\<^bsub>k\<^esub>) t = []\<^sub>l" using lfilter_eq_LNil[of "(\k. \c\\<^bsub>k\<^esub>)"] by blast hence "\ \\<^bsub>c\<^esub>(t) = []\<^sub>l" using proj_def by fastforce thus ?thesis by (simp add: ileI1 lnull_def one_eSuc) qed subsubsection "Projection not Active" lemma proj_not_active[simp]: assumes "enat n < llength t" and "\ \c\\<^bsub>lnth t n\<^esub>" shows "\\<^bsub>c\<^esub>(ltake (Suc n) t) = \\<^bsub>c\<^esub>(ltake n t)" (is "?lhs = ?rhs") proof - from assms have "ltake (enat (Suc n)) t = (ltake (enat n) t) @\<^sub>l ((lnth t n) #\<^sub>l []\<^sub>l)" using ltake_Suc_conv_snoc_lnth by blast hence "?lhs = \\<^bsub>c\<^esub>((ltake (enat n) t) @\<^sub>l ((lnth t n) #\<^sub>l []\<^sub>l))" by simp moreover have "\ = (\\<^bsub>c\<^esub>(ltake (enat n) t)) @\<^sub>l (\\<^bsub>c\<^esub>((lnth t n) #\<^sub>l []\<^sub>l))" by simp moreover from assms have "\\<^bsub>c\<^esub>((lnth t n) #\<^sub>l []\<^sub>l) = []\<^sub>l" by simp ultimately show ?thesis by simp qed lemma proj_not_active_same: assumes "enat n \ (n'::enat)" and "\ lfinite t \ n'-1 < llength t" and "\k. k\n \ k k < llength t \ \c\\<^bsub>lnth t k\<^esub>" shows "\\<^bsub>c\<^esub>(ltake n' t) = \\<^bsub>c\<^esub>(ltake n t)" proof - have "\\<^bsub>c\<^esub>(ltake (n + (n' - n)) t) = \\<^bsub>c\<^esub>((ltake n t) @\<^sub>l (ltake (n'-n) (ldrop n t)))" by (simp add: ltake_plus_conv_lappend) hence "\\<^bsub>c\<^esub>(ltake (n + (n' - n)) t) = (\\<^bsub>c\<^esub>(ltake n t)) @\<^sub>l (\\<^bsub>c\<^esub>(ltake (n'-n) (ldrop n t)))" by simp moreover have "\\<^bsub>c\<^esub>(ltake (n'-n) (ldrop n t)) = []\<^sub>l" proof - have "\k\{lnth (ltake (n' - enat n) (ldrop (enat n) t)) na | na. enat na < llength (ltake (n' - enat n) (ldrop (enat n) t))}. \ \c\\<^bsub>k\<^esub>" proof fix k assume "k\{lnth (ltake (n' - enat n) (ldrop (enat n) t)) na | na. enat na < llength (ltake (n' - enat n) (ldrop (enat n) t))}" then obtain k' where "enat k' < llength (ltake (n' - enat n) (ldrop (enat n) t))" and "k=lnth (ltake (n' - enat n) (ldrop (enat n) t)) k'" by auto have "enat (k' + n) < llength t" proof - from \enat k' < llength (ltake (n' - enat n) (ldrop (enat n) t))\ have "enat k' < n'-n" by simp hence "enat k' + n < n'" using assms(1) enat_min by auto show ?thesis proof cases assume "lfinite t" with \\ lfinite t \ n'-1 < llength t\ have "n'-1 llength t" using eSuc_ile_mono ileI1 by blast with \enat k' + n < n'\ show ?thesis by (simp add: add.commute) next assume "\ lfinite t" hence "llength t = \" using not_lfinite_llength by auto thus ?thesis by simp qed qed moreover have "k = lnth t (k' + n)" proof - from \enat k' < llength (ltake (n' - enat n) (ldrop (enat n) t))\ have "enat k'enat (k' + n) < llength t\ show ?thesis using lnth_ldrop[of n k' t ] using \k = lnth (ltake (n' - enat n) (ldrop (enat n) t)) k'\ by (simp add: add.commute) qed moreover from \enat n \ (n'::enat)\ have "k' + the_enat n\n" by auto moreover from \enat k' < llength (ltake (n' - enat n) (ldrop (enat n) t))\ have "k' + n \c\\<^bsub>k\<^esub>" using \\k. k\n \ k k < llength t \ \c\\<^bsub>lnth t k\<^esub>\ by simp qed hence "\k\lset (ltake (n'-n) (ldrop n t)). \ \c\\<^bsub>k\<^esub>" using lset_conv_lnth[of "(ltake (n' - enat n) (ldrop (enat n) t))"] by simp thus ?thesis using proj_lnull by auto qed moreover from assms have "n + (n' - n) = n'" by (meson enat.distinct(1) enat_add_sub_same enat_diff_cancel_left enat_le_plus_same(1) less_imp_le) ultimately show ?thesis by simp qed subsubsection "Projection Active" lemma proj_active[simp]: assumes "enat i < llength t" "\c\\<^bsub>lnth t i\<^esub>" shows "\\<^bsub>c\<^esub>(ltake (Suc i) t) = (\\<^bsub>c\<^esub>(ltake i t)) @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l []\<^sub>l)" (is "?lhs = ?rhs") proof - from assms have "ltake (enat (Suc i)) t = (ltake (enat i) t) @\<^sub>l ((lnth t i) #\<^sub>l []\<^sub>l)" using ltake_Suc_conv_snoc_lnth by blast hence "?lhs = \\<^bsub>c\<^esub>((ltake (enat i) t) @\<^sub>l ((lnth t i) #\<^sub>l []\<^sub>l))" by simp moreover have "\ = (\\<^bsub>c\<^esub>(ltake (enat i) t)) @\<^sub>l (\\<^bsub>c\<^esub>((lnth t i) #\<^sub>l []\<^sub>l))" by simp moreover from assms have "\\<^bsub>c\<^esub>((lnth t i) #\<^sub>l []\<^sub>l) = (\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l []\<^sub>l" by simp ultimately show ?thesis by simp qed lemma proj_active_append: assumes a1: "(n::nat) \ i" and a2: "enat i < (n'::enat)" and a3: "\ lfinite t \ n'-1 < llength t" and a4: "\c\\<^bsub>lnth t i\<^esub>" and "\i'. (n \ i' \ enat i' i' < llength t \ \c\\<^bsub>lnth t i'\<^esub>) \ (i' = i)" shows "\\<^bsub>c\<^esub>(ltake n' t) = (\\<^bsub>c\<^esub>(ltake n t)) @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l []\<^sub>l)" (is "?lhs = ?rhs") proof - have "?lhs = \\<^bsub>c\<^esub>(ltake (Suc i) t)" proof - from a2 have "Suc i \ n'" by (simp add: Suc_ile_eq) moreover from a3 have "\ lfinite t \ n'-1 < llength t" by simp moreover have "\k. enat k\enat (Suc i) \ k k < llength t \ \c\\<^bsub>lnth t k\<^esub>" proof assume "\k. enat k\enat (Suc i) \ k k < llength t \ \c\\<^bsub>lnth t k\<^esub>" then obtain k where "enat k\enat (Suc i)" and "kc\\<^bsub>lnth t k\<^esub>" by blast moreover from \enat k\enat (Suc i)\ have "enat k\n" using assms by (meson dual_order.trans enat_ord_simps(1) le_SucI) ultimately have "enat k=enat i" using assms using enat_ord_simps(1) by blast with \enat k\enat (Suc i)\ show False by simp qed ultimately show ?thesis using proj_not_active_same[of "Suc i" n' t c] by simp qed also have "\ = (\\<^bsub>c\<^esub>(ltake i t)) @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l []\<^sub>l)" proof - have "i < llength t" proof cases assume "lfinite t" with a3 have "n'-1 < llength t" by simp hence "n' \ llength t" by (metis eSuc_minus_1 enat_minus_mono1 ileI1 not_le) with a2 show "enat i < llength t" by simp next assume "\ lfinite t" thus ?thesis by (metis enat_ord_code(4) llength_eq_infty_conv_lfinite) qed with a4 show ?thesis by simp qed also have "\ = ?rhs" proof - from a1 have "enat n \ enat i" by simp moreover from a2 a3 have "\ lfinite t \ enat i-1 < llength t" using enat_minus_mono1 less_imp_le order.strict_trans1 by blast moreover have "\k. k\n \ enat k enat k < llength t \ \c\\<^bsub>lnth t k\<^esub>" proof assume "\k. k\n \ enat k enat k < llength t \ \c\\<^bsub>lnth t k\<^esub>" then obtain k where "k\n" and "enat kc\\<^bsub>lnth t k\<^esub>" by blast moreover from \enat k have "enat kenat k show False by simp qed ultimately show ?thesis using proj_not_active_same[of n i t c] by simp qed finally show ?thesis by simp qed subsubsection "Same and not Same" lemma proj_same_not_active: assumes "n \ n'" and "enat (n'-1) < llength t" and "\\<^bsub>c\<^esub>(ltake n' t) = \\<^bsub>c\<^esub>(ltake n t)" shows "\k. k\n \ k \c\\<^bsub>lnth t k\<^esub>" proof assume "\k. k\n \ k \c\\<^bsub>lnth t k\<^esub>" then obtain i where "i\n" and "ic\\<^bsub>lnth t i\<^esub>" by blast moreover from \enat (n'-1) and \i have "i\<^bsub>c\<^esub>(ltake (Suc i) t) = (\\<^bsub>c\<^esub>(ltake i t)) @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l []\<^sub>l)" by simp moreover from \i have "Suc i \ n'" by simp hence "lprefix(\\<^bsub>c\<^esub>(ltake (Suc i) t)) (\\<^bsub>c\<^esub>(ltake n' t))" by simp then obtain "tl" where "\\<^bsub>c\<^esub>(ltake n' t) = (\\<^bsub>c\<^esub>(ltake (Suc i) t)) @\<^sub>l tl" using lprefix_conv_lappend by auto moreover from \n\i\ have "lprefix(\\<^bsub>c\<^esub>(ltake n t)) (\\<^bsub>c\<^esub>(ltake i t))" by simp hence "lprefix(\\<^bsub>c\<^esub>(ltake n t)) (\\<^bsub>c\<^esub>(ltake i t))" by simp then obtain "hd" where "\\<^bsub>c\<^esub>(ltake i t) = (\\<^bsub>c\<^esub>(ltake n t)) @\<^sub>l hd" using lprefix_conv_lappend by auto ultimately have "\\<^bsub>c\<^esub>(ltake n' t) = (((\\<^bsub>c\<^esub>(ltake n t)) @\<^sub>l hd) @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l []\<^sub>l)) @\<^sub>l tl" by simp also have "\ = ((\\<^bsub>c\<^esub>(ltake n t)) @\<^sub>l hd) @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l tl)" using lappend_snocL1_conv_LCons2[of "(\\<^bsub>c\<^esub>(ltake n t)) @\<^sub>l hd" "\\<^bsub>c\<^esub>(lnth t i)"] by simp also have "\ = (\\<^bsub>c\<^esub>(ltake n t)) @\<^sub>l (hd @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l tl))" using lappend_assoc by auto also have "\\<^bsub>c\<^esub>(ltake n' t) = (\\<^bsub>c\<^esub>(ltake n' t)) @\<^sub>l []\<^sub>l" by simp finally have "(\\<^bsub>c\<^esub>(ltake n' t)) @\<^sub>l []\<^sub>l = (\\<^bsub>c\<^esub>(ltake n t)) @\<^sub>l (hd @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l tl))" . moreover from assms(3) have "llength (\\<^bsub>c\<^esub>(ltake n' t)) = llength (\\<^bsub>c\<^esub>(ltake n t))" by simp ultimately have "lfinite (\\<^bsub>c\<^esub>(ltake n' t)) \ []\<^sub>l = hd @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l tl)" using assms(3) lappend_eq_lappend_conv[of "\\<^bsub>c\<^esub>(ltake n' t)" "\\<^bsub>c\<^esub>(ltake n t)" "[]\<^sub>l"] by simp moreover have "lfinite (\\<^bsub>c\<^esub>(ltake n' t))" by simp ultimately have "[]\<^sub>l = hd @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l tl)" by simp hence "(\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l tl = []\<^sub>l" using LNil_eq_lappend_iff by auto thus False by simp qed lemma proj_not_same_active: assumes "enat n \ (n'::enat)" and "(\ lfinite t) \ n'-1 < llength t" and "\(\\<^bsub>c\<^esub>(ltake n' t) = \\<^bsub>c\<^esub>(ltake n t))" shows "\k. k\n \ k enat k < llength t \ \c\\<^bsub>lnth t k\<^esub>" proof (rule ccontr) assume "\(\k. k\n \ k enat k < llength t \ \c\\<^bsub>lnth t k\<^esub>)" have "\\<^bsub>c\<^esub>(ltake n' t) = \\<^bsub>c\<^esub>(ltake (enat n) t)" proof cases assume "lfinite t" hence "llength t\\" by (simp add: lfinite_llength_enat) hence "enat (the_enat (llength t)) = llength t" by auto with assms \\ (\k\n. k < n' \ enat k < llength t \ \c\\<^bsub>lnth t k\<^esub>)\ show ?thesis using proj_not_active_same[of n n' t c] by simp next assume "\ lfinite t" with assms \\ (\k\n. k < n' \ enat k < llength t \ \c\\<^bsub>lnth t k\<^esub>)\ show ?thesis using proj_not_active_same[of n n' t c] by simp qed with assms show False by simp qed subsection "Activations" text \ We also introduce an operator to obtain the number of activations of a certain component within a given configuration trace. \ definition nAct :: "'id \ enat \ (cnf llist) \ enat" ("\_ #\<^bsub>_\<^esub>_\") where "\c #\<^bsub>n\<^esub> t\ \ llength (\\<^bsub>c\<^esub>(ltake n t))" lemma nAct_0[simp]: "\c #\<^bsub>0\<^esub> t\ = 0" by (simp add: nAct_def) lemma nAct_NIL[simp]: "\c #\<^bsub>n\<^esub> []\<^sub>l\ = 0" by (simp add: nAct_def) lemma nAct_Null: assumes "llength t \ n" and "\c #\<^bsub>n\<^esub> t\ = 0" shows "\i \c\\<^bsub>lnth t i\<^esub>" proof - from assms have "lnull (\\<^bsub>c\<^esub>(ltake n t))" using nAct_def lnull_def by simp hence "\\<^bsub>c\<^esub>(ltake n t) = []\<^sub>l" using lnull_def by blast hence "(\k\lset (ltake n t). \ \c\\<^bsub>k\<^esub>)" by simp show ?thesis proof (rule ccontr) assume "\ (\i \c\\<^bsub>lnth t i\<^esub>)" then obtain i where "ic\\<^bsub>lnth t i\<^esub>" by blast moreover have "enat i < llength (ltake n t) \ lnth (ltake n t) i = (lnth t i)" proof from \llength t \ n\ have "n = min n (llength t)" using min.orderE by auto hence "llength (ltake n t) = n" by simp with \i show "enat i < llength (ltake n t)" by auto from \i show "lnth (ltake n t) i = (lnth t i)" using lnth_ltake by auto qed hence "(lnth t i \ lset (ltake n t))" using in_lset_conv_lnth[of "lnth t i" "ltake n t"] by blast ultimately show False using \(\k\lset (ltake n t). \ \c\\<^bsub>k\<^esub>)\ by simp qed qed lemma nAct_ge_one[simp]: assumes "llength t \ n" and "i < n" and "\c\\<^bsub>lnth t i\<^esub>" shows "\c #\<^bsub>n\<^esub> t\ \ enat 1" proof (rule ccontr) assume "\ (\c #\<^bsub>n\<^esub> t\ \ enat 1)" hence "\c #\<^bsub>n\<^esub> t\ < enat 1" by simp hence "\c #\<^bsub>n\<^esub> t\ < 1" using enat_1 by simp hence "\c #\<^bsub>n\<^esub> t\ = 0" using Suc_ile_eq \\ enat 1 \ \c #\<^bsub>n\<^esub> t\\ zero_enat_def by auto with \llength t \ n\ have "\i \c\\<^bsub>lnth t i\<^esub>" using nAct_Null by simp with assms show False by simp qed lemma nAct_finite[simp]: assumes "n \ \" shows "\n'. \c #\<^bsub>n\<^esub> t\ = enat n'" proof - from assms have "lfinite (ltake n t)" by simp hence "lfinite (\\<^bsub>c\<^esub>(ltake n t))" by simp hence "\n'. llength (\\<^bsub>c\<^esub>(ltake n t)) = enat n'" using lfinite_llength_enat[of "\\<^bsub>c\<^esub>(ltake n t)"] by simp thus ?thesis using nAct_def by simp qed lemma nAct_enat_the_nat[simp]: assumes "n \ \" shows "enat (the_enat (\c #\<^bsub>n\<^esub> t\)) = \c #\<^bsub>n\<^esub> t\" proof - from assms have "\c #\<^bsub>n\<^esub> t\ \ \" by simp thus ?thesis using enat_the_enat by simp qed subsubsection "Monotonicity and Continuity" lemma nAct_mcont: shows "mcont lSup lprefix Sup (\) (nAct c n)" proof - have "mcont lSup lprefix lSup lprefix (ltake n)" by simp hence "mcont lSup lprefix lSup lprefix (\t. \\<^bsub>c\<^esub>(ltake n t))" using proj_mcont2mcont[of lSup lprefix "(ltake n)"] by simp hence "mcont lSup lprefix Sup (\) (\t. llength (\\<^bsub>c\<^esub>(ltake n t)))" by simp moreover have "nAct c n = (\t. llength (\\<^bsub>c\<^esub>(ltake n t)))" using nAct_def by auto ultimately show ?thesis by simp qed lemma nAct_mono: assumes "n \ n'" shows "\c #\<^bsub>n\<^esub> t\ \ \c #\<^bsub>n'\<^esub> t\" proof - from assms have "lprefix (ltake n t) (ltake n' t)" by simp hence "lprefix (\\<^bsub>c\<^esub>(ltake n t)) (\\<^bsub>c\<^esub>(ltake n' t))" by simp hence "llength (\\<^bsub>c\<^esub>(ltake n t)) \ llength (\\<^bsub>c\<^esub>(ltake n' t))" using lprefix_llength_le[of "(\\<^bsub>c\<^esub>(ltake n t))"] by simp thus ?thesis using nAct_def by simp qed lemma nAct_strict_mono_back: assumes "\c #\<^bsub>n\<^esub> t\ < \c #\<^bsub>n'\<^esub> t\" shows "n < n'" proof (rule ccontr) assume "\ nn'" by simp hence "\c #\<^bsub>n\<^esub> t\ \ \c #\<^bsub>n'\<^esub> t\" using nAct_mono by simp thus False using assms by simp qed subsubsection "Not Active" lemma nAct_not_active[simp]: fixes n::nat and n'::nat and t::"(cnf llist)" and c::'id assumes "enat i < llength t" and "\ \c\\<^bsub>lnth t i\<^esub>" shows "\c #\<^bsub>Suc i\<^esub> t\ = \c #\<^bsub>i\<^esub> t\" proof - from assms have "\\<^bsub>c\<^esub>(ltake (Suc i) t) = \\<^bsub>c\<^esub>(ltake i t)" by simp hence "llength (\\<^bsub>c\<^esub>(ltake (enat (Suc i)) t)) = llength (\\<^bsub>c\<^esub>(ltake i t))" by simp moreover have "llength (\\<^bsub>c\<^esub>(ltake i t)) \ \" using llength_eq_infty_conv_lfinite[of "\\<^bsub>c\<^esub>(ltake (enat i) t)"] by simp ultimately have "llength (\\<^bsub>c\<^esub>(ltake (Suc i) t)) = llength (\\<^bsub>c\<^esub>(ltake i t))" using the_enat_eSuc by simp with nAct_def show ?thesis by simp qed lemma nAct_not_active_same: assumes "enat n \ (n'::enat)" and "n'-1 < llength t" and "\k. enat k\n \ k \c\\<^bsub>lnth t k\<^esub>" shows "\c #\<^bsub>n'\<^esub> t\ = \c #\<^bsub>n\<^esub> t\" using assms proj_not_active_same nAct_def by simp subsubsection "Active" lemma nAct_active[simp]: fixes n::nat and n'::nat and t::"(cnf llist)" and c::'id assumes "enat i < llength t" and "\c\\<^bsub>lnth t i\<^esub>" shows "\c #\<^bsub>Suc i\<^esub> t\ = eSuc (\c #\<^bsub>i\<^esub> t\)" proof - from assms have "\\<^bsub>c\<^esub>(ltake (Suc i) t) = (\\<^bsub>c\<^esub>(ltake i t)) @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l []\<^sub>l)" by simp hence "llength (\\<^bsub>c\<^esub>(ltake (enat (Suc i)) t)) = eSuc (llength (\\<^bsub>c\<^esub>(ltake i t)))" using plus_1_eSuc one_eSuc by simp moreover have "llength (\\<^bsub>c\<^esub>(ltake i t)) \ \" using llength_eq_infty_conv_lfinite[of "\\<^bsub>c\<^esub>(ltake (enat i) t)"] by simp ultimately have "llength (\\<^bsub>c\<^esub>(ltake (Suc i) t)) = eSuc (llength (\\<^bsub>c\<^esub>(ltake i t)))" using the_enat_eSuc by simp with nAct_def show ?thesis by simp qed lemma nAct_active_suc: fixes n::nat and n'::enat and t::"(cnf llist)" and c::'id assumes "\ lfinite t \ n'-1 < llength t" and "n \ i" and "enat i < n'" and "\c\\<^bsub>lnth t i\<^esub>" and "\i'. (n \ i' \ enat i' i' < llength t \ \c\\<^bsub>lnth t i'\<^esub>) \ (i' = i)" shows "\c #\<^bsub>n'\<^esub> t\ = eSuc (\c #\<^bsub>n\<^esub> t\)" proof - from assms have "\\<^bsub>c\<^esub>(ltake n' t) = (\\<^bsub>c\<^esub>(ltake (enat n) t)) @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l []\<^sub>l)" using proj_active_append[of n i n' t c] by blast moreover have "llength ((\\<^bsub>c\<^esub>(ltake (enat n) t)) @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l []\<^sub>l)) = eSuc (llength (\\<^bsub>c\<^esub>(ltake (enat n) t)))" using one_eSuc eSuc_plus_1 by simp ultimately show ?thesis using nAct_def by simp qed lemma nAct_less: assumes "enat k < llength t" and "n \ k" and "k < (n'::enat)" and "\c\\<^bsub>lnth t k\<^esub>" shows "\c #\<^bsub>n\<^esub> t\ < \c #\<^bsub>n'\<^esub> t\" proof - have "\c #\<^bsub>k\<^esub> t\ \ \" by simp then obtain en where en_def: "\c #\<^bsub>k\<^esub> t\ = enat en" by blast moreover have "eSuc (enat en) \ \c #\<^bsub>n'\<^esub> t\" proof - from assms have "Suc k \ n'" using Suc_ile_eq by simp hence "\c #\<^bsub>Suc k\<^esub> t\ \ \c #\<^bsub>n'\<^esub> t\" using nAct_mono by simp moreover from assms have "\c #\<^bsub>Suc k\<^esub> t\ = eSuc (\c #\<^bsub>k\<^esub> t\)" by simp ultimately have "eSuc (\c #\<^bsub>k\<^esub> t\) \ \c #\<^bsub>n'\<^esub> t\" by simp thus ?thesis using en_def by simp qed moreover have "enat en < eSuc (enat en)" by simp ultimately have "enat en < \c #\<^bsub>n'\<^esub> t\" using less_le_trans[of "enat en" "eSuc (enat en)"] by simp moreover have "\c #\<^bsub>n\<^esub> t\ \ enat en" proof - from assms have "\c #\<^bsub>n\<^esub> t\ \ \c #\<^bsub>k\<^esub> t\" using nAct_mono by simp thus ?thesis using en_def by simp qed ultimately show ?thesis using le_less_trans[of "\c #\<^bsub>n\<^esub> t\"] by simp qed lemma nAct_less_active: assumes "n' - 1 < llength t" and "\c #\<^bsub>enat n\<^esub> t\ < \c #\<^bsub>n'\<^esub> t\" shows "\i\n. i \c\\<^bsub>lnth t i\<^esub>" proof (rule ccontr) assume "\ (\i\n. i \c\\<^bsub>lnth t i\<^esub>)" moreover have "enat n \ n'" using assms(2) less_imp_le nAct_strict_mono_back by blast ultimately have "\c #\<^bsub>n\<^esub> t\ = \c #\<^bsub>n'\<^esub> t\" using \n' - 1 < llength t\ nAct_not_active_same by simp thus False using assms by simp qed subsubsection "Same and Not Same" lemma nAct_same_not_active: assumes "\c #\<^bsub>n'\<^esub> inf_llist t\ = \c #\<^bsub>n\<^esub> inf_llist t\" shows "\k\n. k \ \c\\<^bsub>t k\<^esub>" proof (rule ccontr) assume "\(\k\n. k \ \c\\<^bsub>t k\<^esub>)" then obtain k where "k\n" and "kc\\<^bsub>t k\<^esub>" by blast hence "\c #\<^bsub>Suc k\<^esub> inf_llist t\ = eSuc (\c #\<^bsub>k\<^esub> inf_llist t\)" by simp moreover have "\c #\<^bsub>k\<^esub> inf_llist t\\\" by simp ultimately have "\c #\<^bsub>k\<^esub> inf_llist t\ < \c #\<^bsub>Suc k\<^esub> inf_llist t\" by fastforce moreover from \n\k\ have "\c #\<^bsub>n\<^esub> inf_llist t\ \ \c #\<^bsub>k\<^esub> inf_llist t\" using nAct_mono by simp moreover from \k have "Suc k \ n'" by (simp add: Suc_ile_eq) hence "\c #\<^bsub>Suc k\<^esub> inf_llist t\ \ \c #\<^bsub>n'\<^esub> inf_llist t\" using nAct_mono by simp ultimately show False using assms by simp qed lemma nAct_not_same_active: assumes "\c #\<^bsub>enat n\<^esub> t\ < \c #\<^bsub>n'\<^esub> t\" and "\ lfinite t \ n' - 1 < llength t" shows "\(i::nat)\n. enat i< n' \ i \c\\<^bsub>lnth t i\<^esub>" proof - from assms have "llength(\\<^bsub>c\<^esub>(ltake n t)) < llength (\\<^bsub>c\<^esub>(ltake n' t))" using nAct_def by simp hence "\\<^bsub>c\<^esub>(ltake n' t) \ \\<^bsub>c\<^esub>(ltake n t)" by auto moreover from assms have "enat n < n'" using nAct_strict_mono_back[of c "enat n" t n'] by simp ultimately show ?thesis using proj_not_same_active[of n n' t c] assms by simp qed lemma nAct_less_llength_active: assumes "x < llength (\\<^bsub>c\<^esub>(t))" and "enat x = \c #\<^bsub>enat n'\<^esub> t\" shows "\(i::nat)\n'. i \c\\<^bsub>lnth t i\<^esub>" proof - have "llength(\\<^bsub>c\<^esub>(ltake n' t)) < llength (\\<^bsub>c\<^esub>(t))" using assms(1) assms(2) nAct_def by auto hence "llength(\\<^bsub>c\<^esub>(ltake n' t)) < llength (\\<^bsub>c\<^esub>(ltake (llength t) t))" by (simp add: ltake_all) hence "\c #\<^bsub>enat n'\<^esub> t\ < \c #\<^bsub>llength t\<^esub> t\" using nAct_def by simp moreover have "\ lfinite t \ llength t - 1 < llength t" proof (rule Meson.imp_to_disjD[OF HOL.impI]) assume "lfinite t" hence "llength t \ \" by (simp add: llength_eq_infty_conv_lfinite) moreover have "llength t>0" proof - from \x < llength (\\<^bsub>c\<^esub>(t))\ have "llength (\\<^bsub>c\<^esub>(t))>0" by auto thus ?thesis using proj_llength order.strict_trans2 by blast qed ultimately show "llength t - 1 < llength t" by (metis One_nat_def \lfinite t\ diff_Suc_less enat_ord_simps(2) idiff_enat_enat lfinite_conv_llength_enat one_enat_def zero_enat_def) qed ultimately show ?thesis using nAct_not_same_active[of c n' t "llength t"] by simp qed lemma nAct_exists: assumes "x < llength (\\<^bsub>c\<^esub>(t))" shows "\(n'::nat). enat x = \c #\<^bsub>n'\<^esub> t\" proof - have "x < llength (\\<^bsub>c\<^esub>(t)) \ (\(n'::nat). enat x = \c #\<^bsub>n'\<^esub> t\)" proof (induction x) case 0 thus ?case by (metis nAct_0 zero_enat_def) next case (Suc x) show ?case proof assume "Suc x < llength (\\<^bsub>c\<^esub>(t))" hence "x < llength (\\<^bsub>c\<^esub>(t))" using Suc_ile_eq less_imp_le by auto with Suc.IH obtain n' where "enat x = \c #\<^bsub>enat n'\<^esub> t\" by blast with \x < llength (\\<^bsub>c\<^esub>(t))\ have "\i\n'. i < llength t \ \c\\<^bsub>lnth t i\<^esub>" using nAct_less_llength_active[of x c t n'] by simp then obtain i where "i\n'" and "ic\\<^bsub>lnth t i\<^esub>" and "\k. n'\k \ k k \c\\<^bsub>lnth t k\<^esub>" using lActive_least[of n' t c] by auto moreover from \i have "\ lfinite t \ enat (Suc i) - 1 < llength t" by (simp add: one_enat_def) moreover have "enat i < enat (Suc i)" by simp moreover have "\i'. (n' \ i' \ enat i' i' \c\\<^bsub>lnth t i'\<^esub>) \ (i' = i)" proof (rule HOL.impI[THEN HOL.allI]) fix i' assume "n' \ i' \ enat i' i' \c\\<^bsub>lnth t i'\<^esub>" with \\k. n'\k \ k k \c\\<^bsub>lnth t k\<^esub>\ show "i'=i" by fastforce qed ultimately have "\c #\<^bsub>Suc i\<^esub> t\ = eSuc (\c #\<^bsub>n'\<^esub> t\)" using nAct_active_suc[of t "Suc i" n' i c] by simp with \enat x = \c #\<^bsub>enat n'\<^esub> t\\ have "\c #\<^bsub>Suc i\<^esub> t\ = eSuc (enat x)" by simp thus "\n'. enat (Suc x) = \c #\<^bsub>enat n'\<^esub> t\" by (metis eSuc_enat) qed qed with assms show ?thesis by simp qed subsection "Projection and Activation" text \ In the following we provide some properties about the relationship between the projection and activations operator. \ lemma nAct_le_proj: "\c #\<^bsub>n\<^esub> t\ \ llength (\\<^bsub>c\<^esub>(t))" proof - from nAct_def have "\c #\<^bsub>n\<^esub> t\ = llength (\\<^bsub>c\<^esub>(ltake n t))" by simp moreover have "llength (\\<^bsub>c\<^esub>(ltake n t)) \ llength (\\<^bsub>c\<^esub>(t))" proof - have "lprefix (ltake n t) t" by simp hence "lprefix (\\<^bsub>c\<^esub>(ltake n t)) (\\<^bsub>c\<^esub>(t))" by simp hence "llength (\\<^bsub>c\<^esub>(ltake n t)) \ llength (\\<^bsub>c\<^esub>(t))" using lprefix_llength_le by blast thus ?thesis by auto qed thus ?thesis using nAct_def by simp qed lemma proj_nAct: assumes "(enat n < llength t)" shows "\\<^bsub>c\<^esub>(ltake n t) = ltake (\c #\<^bsub>n\<^esub> t\) (\\<^bsub>c\<^esub>(t))" (is "?lhs = ?rhs") proof - have "?lhs = ltake (llength (\\<^bsub>c\<^esub>(ltake n t))) (\\<^bsub>c\<^esub>(ltake n t))" using ltake_all[of "\\<^bsub>c\<^esub>(ltake n t)" "llength (\\<^bsub>c\<^esub>(ltake n t))"] by simp also have "\ = ltake (llength (\\<^bsub>c\<^esub>(ltake n t))) ((\\<^bsub>c\<^esub>(ltake n t)) @\<^sub>l (\\<^bsub>c\<^esub>(ldrop n t)))" using ltake_lappend1[of "llength (\\<^bsub>c\<^esub>(ltake (enat n) t))" "\\<^bsub>c\<^esub>(ltake n t)" "(\\<^bsub>c\<^esub>(ldrop n t))"] by simp also have "\ = ltake (\c #\<^bsub>n\<^esub> t\) ((\\<^bsub>c\<^esub>(ltake n t)) @\<^sub>l (\\<^bsub>c\<^esub>(ldrop n t)))" using nAct_def by simp also have "\ = ltake (\c #\<^bsub>n\<^esub> t\) (\\<^bsub>c\<^esub>((ltake (enat n) t) @\<^sub>l (ldrop n t)))" by simp also have "\ = ltake (\c #\<^bsub>n\<^esub> t\) (\\<^bsub>c\<^esub>(t))" using lappend_ltake_ldrop[of n t] by simp finally show ?thesis by simp qed lemma proj_active_nth: assumes "enat (Suc i) < llength t" "\c\\<^bsub>lnth t i\<^esub>" shows "lnth (\\<^bsub>c\<^esub>(t)) (the_enat (\c #\<^bsub>i\<^esub> t\)) = \\<^bsub>c\<^esub>(lnth t i)" proof - from assms have "enat i < llength t" using Suc_ile_eq[of i "llength t"] by auto with assms have "\\<^bsub>c\<^esub>(ltake (Suc i) t) = (\\<^bsub>c\<^esub>(ltake i t)) @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l []\<^sub>l)" by simp moreover have "lnth ((\\<^bsub>c\<^esub>(ltake i t)) @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l []\<^sub>l)) (the_enat (llength (\\<^bsub>c\<^esub>(ltake i t)))) = \\<^bsub>c\<^esub>(lnth t i)" proof - have "\ lnull ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l []\<^sub>l)" by simp moreover have "lfinite (\\<^bsub>c\<^esub>(ltake i t))" by simp ultimately have "lnth ((\\<^bsub>c\<^esub>(ltake i t)) @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l []\<^sub>l)) (the_enat (llength (\\<^bsub>c\<^esub>(ltake i t)))) = lhd ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l []\<^sub>l)" by simp also have "\ = \\<^bsub>c\<^esub>(lnth t i)" by simp finally show "lnth ((\\<^bsub>c\<^esub>(ltake i t)) @\<^sub>l ((\\<^bsub>c\<^esub>(lnth t i)) #\<^sub>l []\<^sub>l)) (the_enat (llength (\\<^bsub>c\<^esub>(ltake i t)))) = \\<^bsub>c\<^esub>(lnth t i)" by simp qed ultimately have "\\<^bsub>c\<^esub>(lnth t i) = lnth (\\<^bsub>c\<^esub>(ltake (Suc i) t)) (the_enat (llength (\\<^bsub>c\<^esub>(ltake i t))))" by simp also have "\ = lnth (\\<^bsub>c\<^esub>(ltake (Suc i) t)) (the_enat (\c #\<^bsub>i\<^esub> t\))" using nAct_def by simp also have "\ = lnth (ltake (\c #\<^bsub>Suc i\<^esub> t\) (\\<^bsub>c\<^esub>(t))) (the_enat (\c #\<^bsub>i\<^esub> t\))" using proj_nAct[of "Suc i" t c] assms by simp also have "\ = lnth (\\<^bsub>c\<^esub>(t)) (the_enat (\c #\<^bsub>i\<^esub> t\))" proof - from assms have "\c #\<^bsub>Suc i\<^esub> t\ = eSuc (\c #\<^bsub>i\<^esub> t\)" using \enat i < llength t\ by simp moreover have "\c #\<^bsub>i\<^esub> t\ < eSuc (\c #\<^bsub>i\<^esub> t\)" using iless_Suc_eq[of "the_enat (\c #\<^bsub>enat i\<^esub> t\)"] by simp ultimately have "\c #\<^bsub>i\<^esub> t\ < (\c #\<^bsub>Suc i\<^esub> t\)" by simp hence "enat (the_enat (\c #\<^bsub>Suc i\<^esub> t\)) > enat (the_enat (\c #\<^bsub>i\<^esub> t\))" by simp thus ?thesis using lnth_ltake[of "the_enat (\c #\<^bsub>i\<^esub> t\)" "the_enat (\c #\<^bsub>enat (Suc i)\<^esub> t\)" "\\<^bsub>c\<^esub>(t)"] by simp qed finally show ?thesis .. qed lemma nAct_eq_proj: assumes "\(\i\n. \c\\<^bsub>lnth t i\<^esub>)" shows "\c #\<^bsub>n\<^esub> t\ = llength (\\<^bsub>c\<^esub>(t))" (is "?lhs = ?rhs") proof - from nAct_def have "?lhs = llength (\\<^bsub>c\<^esub>(ltake n t))" by simp moreover from assms have "\(n'::nat)\llength t. n'\n \ (\ \c\\<^bsub>lnth t n'\<^esub>)" by simp hence "\\<^bsub>c\<^esub>(t) = \\<^bsub>c\<^esub>(ltake n t)" using proj_ltake by simp ultimately show ?thesis by simp qed lemma nAct_llength_proj: assumes "\i\n. \c\\<^bsub>t i\<^esub>" shows "llength (\\<^bsub>c\<^esub>(inf_llist t)) \ eSuc (\c #\<^bsub>n\<^esub> inf_llist t\)" proof - from \\i\n. \c\\<^bsub>t i\<^esub>\ obtain i where "i\n" and "\c\\<^bsub>t i\<^esub>" and "\ (\k\n. k < i \ k < llength (inf_llist t) \ \c\\<^bsub>t k\<^esub>)" using lActive_least[of n "inf_llist t" c] by auto moreover have "llength (\\<^bsub>c\<^esub>(inf_llist t)) \ \c #\<^bsub>Suc i\<^esub> inf_llist t\" using nAct_le_proj by simp moreover have "eSuc (\c #\<^bsub>n\<^esub> inf_llist t\) = \c #\<^bsub>Suc i\<^esub> inf_llist t\" proof - have "enat (Suc i) < llength (inf_llist t)" by simp moreover have "i < Suc i" by simp moreover from \\ (\k\n. k < i \ k < llength (inf_llist t) \ \c\\<^bsub>t k\<^esub>)\ have "\i'. n \ i' \ i' < Suc i \ \c\\<^bsub>lnth (inf_llist t) i'\<^esub> \ i' = i" by fastforce ultimately show ?thesis using nAct_active_suc \i\n\ \\c\\<^bsub>t i\<^esub>\ by simp qed ultimately show ?thesis by simp qed subsection "Least not Active" text \ In the following, we introduce an operator to obtain the least point in time before a certain point in time where a component was deactivated. \ definition lNAct :: "'id \ (nat \ cnf) \ nat \ nat" ("\_ \ _\\<^bsub>_\<^esub>") where "\c \ t\\<^bsub>n\<^esub> \ (LEAST n'. n=n' \ (n' (\k. k\n' \ k \c\\<^bsub>t k\<^esub>)))" lemma lNact0[simp]: "\c \ t\\<^bsub>0\<^esub> = 0" by (simp add: lNAct_def) lemma lNact_least: assumes "n=n' \ n' (\k. k\n' \ k \c\\<^bsub>t k\<^esub>)" shows "\c \ t\\<^bsub>n\<^esub> \ n'" using Least_le[of "\n'. n=n' \ (n' (\k. k\n' \ k \c\\<^bsub>t k\<^esub>))" n'] lNAct_def using assms by auto lemma lNAct_ex: "\c \ t\\<^bsub>n\<^esub>=n \ \c \ t\\<^bsub>n\<^esub> (\k. k\\c \ t\\<^bsub>n\<^esub> \ k \c\\<^bsub>t k\<^esub>)" proof - let ?P="\n'. n=n' \ n' (\k. k\n' \ k \c\\<^bsub>t k\<^esub>)" from lNAct_def have "\c \ t\\<^bsub>n\<^esub> = (LEAST n'. ?P n')" by simp moreover have "?P n" by simp with LeastI have "?P (LEAST n'. ?P n')" . ultimately show ?thesis by auto qed lemma lNact_notActive: fixes c t n k assumes "k\\c \ t\\<^bsub>n\<^esub>" and "k\c\\<^bsub>t k\<^esub>" by (metis assms lNAct_ex leD) lemma lNactGe: fixes c t n n' assumes "n' \ \c \ t\\<^bsub>n\<^esub>" and "\c\\<^bsub>t n'\<^esub>" shows "n' \ n" using assms lNact_notActive leI by blast lemma lNactLe[simp]: fixes n n' shows "\c \ t\\<^bsub>n\<^esub> \ n" using lNAct_ex less_or_eq_imp_le by blast lemma lNactLe_nact: fixes n n' assumes "n'=n \ (n' (\k. k\n' \ k \c\\<^bsub>t k\<^esub>))" shows "\c \ t\\<^bsub>n\<^esub> \ n'" using assms lNAct_def Least_le[of "\n'. n=n' \ (n' (\k. k\n' \ k \c\\<^bsub>t k\<^esub>))"] by auto lemma lNact_active: fixes cid t n assumes "\kcid\\<^bsub>t k\<^esub>" shows "\cid \ t\\<^bsub>n\<^esub> = n" using assms lNAct_ex by blast lemma nAct_mono_back: fixes c t and n and n' assumes "\c #\<^bsub>n'\<^esub> inf_llist t\ \ \c #\<^bsub>n\<^esub> inf_llist t\" shows "n'\\c \ t\\<^bsub>n\<^esub>" proof cases assume "\c #\<^bsub>n'\<^esub> inf_llist t\ = \c #\<^bsub>n\<^esub> inf_llist t\" thus ?thesis proof cases assume "n'\n" - thus ?thesis using lNactLe by (metis HOL.no_atp(11)) + thus ?thesis + by (rule order_trans[OF lNactLe]) next assume "\ n'\n" hence "n'\c #\<^bsub>n'\<^esub> inf_llist t\ = \c #\<^bsub>n\<^esub> inf_llist t\\ have "\k. k\n' \ k \c\\<^bsub>t k\<^esub>" by (metis enat_ord_simps(1) enat_ord_simps(2) nAct_same_not_active) thus ?thesis using lNactLe_nact by (simp add: \n' < n\) qed next assume "\\c #\<^bsub>n'\<^esub> inf_llist t\ = \c #\<^bsub>n\<^esub> inf_llist t\" with assms have "\c #\<^bsub>enat n'\<^esub> inf_llist t\ > \c #\<^bsub>enat n\<^esub> inf_llist t\" by simp hence "n' > n" using nAct_strict_mono_back[of c "enat n" "inf_llist t" "enat n'"] by simp thus ?thesis by (meson dual_order.strict_implies_order lNactLe le_trans) qed lemma nAct_mono_lNact: assumes "\c \ t\\<^bsub>n\<^esub> \ n'" shows "\c #\<^bsub>n\<^esub> inf_llist t\ \ \c #\<^bsub>n'\<^esub> inf_llist t\" proof - have "\k. k\\c \ t\\<^bsub>n\<^esub> \ k \c\\<^bsub>t k\<^esub>" using lNact_notActive by auto moreover have "enat n - 1 < llength (inf_llist t)" by (simp add: one_enat_def) moreover from \\c \ t\\<^bsub>n\<^esub> \ n'\ have "enat \c \ t\\<^bsub>n\<^esub> \ enat n" by simp ultimately have "\c #\<^bsub>n\<^esub> inf_llist t\=\c #\<^bsub>\c \ t\\<^bsub>n\<^esub>\<^esub> inf_llist t\" using nAct_not_active_same by simp thus ?thesis using nAct_mono assms by simp qed subsection "Next Active" text \ In the following, we introduce an operator to obtain the next point in time when a component is activated. \ definition nxtAct :: "'id \ (nat \ cnf) \ nat \ nat" ("\_ \ _\\<^bsub>_\<^esub>") where "\c \ t\\<^bsub>n\<^esub> \ (THE n'. n'\n \ \c\\<^bsub>t n'\<^esub> \ (\k. k\n \ k \c\\<^bsub>t k\<^esub>))" lemma nxtActI: fixes n::nat and t::"nat \ cnf" and c::'id assumes "\i\n. \c\\<^bsub>t i\<^esub>" shows "\c \ t\\<^bsub>n\<^esub> \ n \ \c\\<^bsub>t \c \ t\\<^bsub>n\<^esub>\<^esub> \ (\k. k\n \ k<\c \ t\\<^bsub>n\<^esub> \ \c\\<^bsub>t k\<^esub>)" proof - let ?P = "THE n'. n'\n \ \c\\<^bsub>t n'\<^esub> \ (\k. k\n \ k \c\\<^bsub>t k\<^esub>)" from assms obtain i where "i\n \ \c\\<^bsub>t i\<^esub> \ (\k. k\n \ k \c\\<^bsub>t k\<^esub>)" using lActive_least[of n "inf_llist t" c] by auto moreover have "(\x. n \ x \ \c\\<^bsub>t x\<^esub> \ \ (\k\n. k < x \ \c\\<^bsub>t k\<^esub>) \ x = i)" proof - fix x assume "n \ x \ \c\\<^bsub>t x\<^esub> \ \ (\k\n. k < x \ \c\\<^bsub>t k\<^esub>)" show "x = i" proof (rule ccontr) assume "\ (x = i)" thus False using \i\n \ \c\\<^bsub>t i\<^esub> \ (\k. k\n \ k \c\\<^bsub>t k\<^esub>)\ \n \ x \ \c\\<^bsub>t x\<^esub> \ \ (\k\n. k < x \ \c\\<^bsub>t k\<^esub>)\ by fastforce qed qed ultimately have "(?P) \ n \ \c\\<^bsub>t (?P)\<^esub> \ (\k. k\n \ k \c\\<^bsub>t k\<^esub>)" using theI[of "\n'. n'\n \ \c\\<^bsub>t n'\<^esub> \ (\k. k\n \ k \c\\<^bsub>t k\<^esub>)"] by blast thus ?thesis using nxtAct_def[of c t n] by metis qed lemma nxtActLe: fixes n n' assumes "\i\n. \c\\<^bsub>t i\<^esub>" shows "n \ \c \ t\\<^bsub>n\<^esub>" by (simp add: assms nxtActI) lemma nxtAct_eq: assumes "n'\n" and "\c\\<^bsub>t n'\<^esub>" and "\n''\n. n'' \ \c\\<^bsub>t n''\<^esub>" shows "n' = \c \ t\\<^bsub>n\<^esub>" by (metis assms(1) assms(2) assms(3) nxtActI linorder_neqE_nat nxtActLe) lemma nxtAct_active: fixes i::nat and t::"nat \ cnf" and c::'id assumes "\c\\<^bsub>t i\<^esub>" shows "\c \ t\\<^bsub>i\<^esub> = i" by (metis assms le_eq_less_or_eq nxtActI) lemma nxtActive_no_active: assumes "\!i. i\n \ \c\\<^bsub>t i\<^esub>" shows "\ (\i'\Suc \c \ t\\<^bsub>n\<^esub>. \c\\<^bsub>t i'\<^esub>)" proof assume "\i'\Suc \c \ t\\<^bsub>n\<^esub>. \c\\<^bsub>t i'\<^esub>" then obtain i' where "i'\Suc \c \ t\\<^bsub>n\<^esub>" and "\c\\<^bsub>t i'\<^esub>" by auto moreover from assms(1) have "\c \ t\\<^bsub>n\<^esub>\n" using nxtActI by auto ultimately have "i'\n" and "\c\\<^bsub>t i'\<^esub>" and "i'\\c \ t\\<^bsub>n\<^esub>" by auto moreover from assms(1) have "\c\\<^bsub>t \c \ t\\<^bsub>n\<^esub>\<^esub>" and "\c \ t\\<^bsub>n\<^esub>\n" using nxtActI by auto ultimately show False using assms(1) by auto qed lemma nxt_geq_lNact[simp]: assumes "\i\n. \c\\<^bsub>t i\<^esub>" shows "\c \ t\\<^bsub>n\<^esub>\\c \ t\\<^bsub>n\<^esub>" proof - from assms have "n \ \c \ t\\<^bsub>n\<^esub>" using nxtActLe by simp moreover have "\c \ t\\<^bsub>n\<^esub>\n" by simp ultimately show ?thesis by arith qed lemma active_geq_nxtAct: assumes "\c\\<^bsub>t i\<^esub>" and "the_enat (\c #\<^bsub>i\<^esub> inf_llist t\)\the_enat (\c #\<^bsub>n\<^esub> inf_llist t\)" shows "i\\c \ t\\<^bsub>n\<^esub>" proof cases assume "\c #\<^bsub>i\<^esub> inf_llist t\=\c #\<^bsub>n\<^esub> inf_llist t\" show ?thesis proof (rule ccontr) assume "\ i\\c \ t\\<^bsub>n\<^esub>" hence "i<\c \ t\\<^bsub>n\<^esub>" by simp with \\c #\<^bsub>i\<^esub> inf_llist t\=\c #\<^bsub>n\<^esub> inf_llist t\\ have "\ (\k\i. k < n \ \c\\<^bsub>t k\<^esub>)" by (metis enat_ord_simps(1) leD leI nAct_same_not_active) moreover have "\ (\k\n. k <\c \ t\\<^bsub>n\<^esub> \ \c\\<^bsub>t k\<^esub>)" using nxtActI by blast ultimately have "\ (\k\i. k <\c \ t\\<^bsub>n\<^esub> \ \c\\<^bsub>t k\<^esub>)" by auto with \i<\c \ t\\<^bsub>n\<^esub>\ show False using \\c\\<^bsub>t i\<^esub>\ by simp qed next assume "\\c #\<^bsub>i\<^esub> inf_llist t\=\c #\<^bsub>n\<^esub> inf_llist t\" moreover from \the_enat (\c #\<^bsub>i\<^esub> inf_llist t\)\the_enat (\c #\<^bsub>n\<^esub> inf_llist t\)\ have "\c #\<^bsub>i\<^esub> inf_llist t\\\c #\<^bsub>n\<^esub> inf_llist t\" by (metis enat.distinct(2) enat_ord_simps(1) nAct_enat_the_nat) ultimately have "\c #\<^bsub>i\<^esub> inf_llist t\>\c #\<^bsub>n\<^esub> inf_llist t\" by simp hence "i>n" using nAct_strict_mono_back[of c n "inf_llist t" i] by simp with \\c\\<^bsub>t i\<^esub>\ show ?thesis by (meson dual_order.strict_implies_order leI nxtActI) qed lemma nAct_same: assumes "\c \ t\\<^bsub>n\<^esub> \ n'" and "n' \ \c \ t\\<^bsub>n\<^esub>" shows "the_enat (\c #\<^bsub>enat n'\<^esub> inf_llist t\) = the_enat (\c #\<^bsub>enat n\<^esub> inf_llist t\)" proof cases assume "n \ n'" moreover have "n' - 1 < llength (inf_llist t)" by simp moreover have "\ (\i\n. i \c\\<^bsub>t i\<^esub>)" by (meson assms(2) less_le_trans nxtActI) ultimately show ?thesis using nAct_not_active_same by (simp add: one_enat_def) next assume "\ n \ n'" hence "n' < n" by simp moreover have "n - 1 < llength (inf_llist t)" by simp moreover have "\ (\i\n'. i < n \ \c\\<^bsub>t i\<^esub>)" by (metis \\ n \ n'\ assms(1) dual_order.trans lNAct_ex) ultimately show ?thesis using nAct_not_active_same[of n' n] by (simp add: one_enat_def) qed lemma nAct_mono_nxtAct: assumes "\i\n. \c\\<^bsub>t i\<^esub>" and "\c \ t\\<^bsub>n\<^esub> \ n'" shows "\c #\<^bsub>n\<^esub> inf_llist t\ \ \c #\<^bsub>n'\<^esub> inf_llist t\" proof - from assms have "\c #\<^bsub>\c \ t\\<^bsub>n\<^esub>\<^esub> inf_llist t\ \ \c #\<^bsub>n'\<^esub> inf_llist t\" using nAct_mono assms by simp moreover have "\c #\<^bsub>\c \ t\\<^bsub>n\<^esub>\<^esub> inf_llist t\=\c #\<^bsub>n\<^esub> inf_llist t\" proof - from assms have "\k. k\n \ k<\c \ t\\<^bsub>n\<^esub> \ \c\\<^bsub>t k\<^esub>" and "n \ \c \ t\\<^bsub>n\<^esub>" using nxtActI by auto moreover have "enat \c \ t\\<^bsub>n\<^esub> - 1 < llength (inf_llist t)" by (simp add: one_enat_def) ultimately show ?thesis using nAct_not_active_same[of n "\c \ t\\<^bsub>n\<^esub>"] by auto qed ultimately show ?thesis by simp qed subsection "Latest Activation" text \ In the following, we introduce an operator to obtain the latest point in time when a component is activated. \ abbreviation latestAct_cond:: "'id \ trace \ nat \ nat \ bool" where "latestAct_cond c t n n' \ n' \c\\<^bsub>t n'\<^esub>" definition latestAct:: "'id \ trace \ nat \ nat" ("\_ \ _\\<^bsub>_\<^esub>") where "latestAct c t n = (GREATEST n'. latestAct_cond c t n n')" lemma latestActEx: assumes "\n'nid\\<^bsub>t n'\<^esub>" shows "\n'. latestAct_cond nid t n n' \ (\n''. latestAct_cond nid t n n'' \ n'' \ n')" proof - from assms obtain n' where "latestAct_cond nid t n n'" by auto moreover have "\n''>n. \ latestAct_cond nid t n n''" by simp ultimately obtain n' where "latestAct_cond nid t n n' \ (\n''. latestAct_cond nid t n n'' \ n'' \ n')" using boundedGreatest[of "latestAct_cond nid t n" n'] by blast thus ?thesis .. qed lemma latestAct_prop: assumes "\n'nid\\<^bsub>t n'\<^esub>" shows "\nid\\<^bsub>t (latestAct nid t n)\<^esub>" and "latestAct nid t nnid\\<^bsub>t \nid \ t\\<^bsub>n\<^esub>\<^esub>" and "latestAct nid t n \nid \ t\\<^bsub>n\<^esub>" proof - from assms latestActEx have "n' \ (GREATEST x. latestAct_cond nid t n x)" using Greatest_le_nat[of "latestAct_cond nid t n"] by blast thus ?thesis using latestAct_def by auto qed lemma latestActNxt: assumes "\n'nid\\<^bsub>t n'\<^esub>" shows "\nid \ t\\<^bsub>\nid \ t\\<^bsub>n\<^esub>\<^esub>=\nid \ t\\<^bsub>n\<^esub>" using assms latestAct_prop(1) nxtAct_active by auto lemma latestActNxtAct: assumes "\n'\n. \tid\\<^bsub>t n'\<^esub>" and "\n'tid\\<^bsub>t n'\<^esub>" shows "\tid \ t\\<^bsub>n\<^esub> > \tid \ t\\<^bsub>n\<^esub>" by (meson assms latestAct_prop(2) less_le_trans nxtActI zero_le) lemma latestActless: assumes "\n'\n\<^sub>s. n' \nid\\<^bsub>t n'\<^esub>" shows "\nid \ t\\<^bsub>n\<^esub>\n\<^sub>s" by (meson assms dual_order.trans latestAct_less) lemma latestActEq: fixes nid::'id assumes "\nid\\<^bsub>t n'\<^esub>" and "\(\n''>n'. n'' \nid\\<^bsub>t n'\<^esub>)" and "n'nid \ t\\<^bsub>n\<^esub> = n'" using latestAct_def proof have "(GREATEST n'. latestAct_cond nid t n n') = n'" proof (rule Greatest_equality[of "latestAct_cond nid t n" n']) from assms(1) assms (3) show "latestAct_cond nid t n n'" by simp next fix y assume "latestAct_cond nid t n y" hence "\nid\\<^bsub>t y\<^esub>" and "y n'" using assms(1) assms (2) leI by blast qed thus "n' = (GREATEST n'. latestAct_cond nid t n n')" by simp qed subsection "Last Activation" text \ In the following we introduce an operator to obtain the latest point in time where a certain component was activated within a certain configuration trace. \ definition lActive :: "'id \ (nat \ cnf) \ nat" ("\_ \ _\") where "\c \ t\ \ (GREATEST i. \c\\<^bsub>t i\<^esub>)" lemma lActive_active: assumes "\c\\<^bsub>t i\<^esub>" and "\n' > n. \ (\c\\<^bsub>t n'\<^esub>)" shows "\c\\<^bsub>t (\c \ t\)\<^esub>" proof - from assms obtain i' where "\c\\<^bsub>t i'\<^esub>" and "\y. \c\\<^bsub>t y\<^esub> \ y \ i'" using boundedGreatest[of "\i'. \c\\<^bsub>t i'\<^esub>" i n] by blast thus ?thesis using lActive_def Nat.GreatestI_nat[of "\i'. \c\\<^bsub>t i'\<^esub>"] by metis qed lemma lActive_less: assumes "\c\\<^bsub>t i\<^esub>" and "\n' > n. \ (\c\\<^bsub>t n'\<^esub>)" shows "\c \ t\ \ n" proof (rule ccontr) assume "\ \c \ t\ \ n" hence "\c \ t\ > n" by simp moreover from assms have "\c\\<^bsub>t (\c \ t\)\<^esub>" using lActive_active by simp ultimately show False using assms by simp qed lemma lActive_greatest: assumes "\c\\<^bsub>t i\<^esub>" and "\n' > n. \ (\c\\<^bsub>t n'\<^esub>)" shows "i \ \c \ t\" proof - from assms obtain i' where "\c\\<^bsub>t i'\<^esub>" and "\y. \c\\<^bsub>t y\<^esub> \ y \ i'" using boundedGreatest[of "\i'. \c\\<^bsub>t i'\<^esub>" i n] by blast with assms show ?thesis using lActive_def Nat.Greatest_le_nat[of "\i'. \c\\<^bsub>t i'\<^esub>" i] by metis qed lemma lActive_greater_active: assumes "n > \c \ t\" and "\n'' > n'. \ \c\\<^bsub>t n''\<^esub>" shows "\ \c\\<^bsub>t n\<^esub>" proof (rule ccontr) assume "\ \ \c\\<^bsub>t n\<^esub>" with \\n'' > n'. \ \c\\<^bsub>t n''\<^esub>\ have "n \ \c \ t\" using lActive_greatest by simp thus False using assms by simp qed lemma lActive_greater_active_all: assumes "\n'' > n'. \ \c\\<^bsub>t n''\<^esub>" shows "\(\n > \c \ t\. \c\\<^bsub>t n\<^esub>)" proof (rule ccontr) assume "\\(\n > \c \ t\. \c\\<^bsub>t n\<^esub>)" then obtain "n" where "n>\c \ t\" and "\c\\<^bsub>t n\<^esub>" by blast with \\n'' > n'. \ (\c\\<^bsub>t n''\<^esub>)\ have "\ \c\\<^bsub>t n\<^esub>" using lActive_greater_active by simp with \\c\\<^bsub>t n\<^esub>\ show False by simp qed lemma lActive_equality: assumes "\c\\<^bsub>t i\<^esub>" and "(\x. \c\\<^bsub>t x\<^esub> \ x \ i)" shows "\c \ t\ = i" unfolding lActive_def using assms Greatest_equality[of "\i'. \c\\<^bsub>t i'\<^esub>"] by simp lemma nxtActive_lactive: assumes "\i\n. \c\\<^bsub>t i\<^esub>" and "\ (\i>\c \ t\\<^bsub>n\<^esub>. \c\\<^bsub>t i\<^esub>)" shows "\c \ t\\<^bsub>n\<^esub>=\c \ t\" proof - from assms(1) have "\c\\<^bsub>t \c \ t\\<^bsub>n\<^esub>\<^esub>" using nxtActI by auto moreover from assms have "\ (\i'\Suc \c \ t\\<^bsub>n\<^esub>. \c\\<^bsub>t i'\<^esub>)" using nxtActive_no_active by simp hence "(\x. \c\\<^bsub>t x\<^esub> \ x \ \c \ t\\<^bsub>n\<^esub>)" using not_less_eq_eq by auto ultimately show ?thesis using \\ (\i'\Suc \c \ t\\<^bsub>n\<^esub>. \c\\<^bsub>t i'\<^esub>)\ lActive_equality by simp qed subsection "Mapping Time Points" text \ In the following we introduce two operators to map time-points between configuration traces and behavior traces. \ subsubsection "Configuration Trace to Behavior Trace" text \ First we provide an operator which maps a point in time of a configuration trace to the corresponding point in time of a behavior trace. \ definition cnf2bhv :: "'id \ (nat \ cnf) \ nat \ nat" ("\<^bsub>_\<^esub>\\<^bsub>_\<^esub>(_)" [150,150,150] 110) where "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n) \ the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1 + (n - \c \ t\)" lemma cnf2bhv_mono: assumes "n'\n" shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n') \ \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n)" by (simp add: assms cnf2bhv_def diff_le_mono) lemma cnf2bhv_mono_strict: assumes "n\\c \ t\" and "n'>n" shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n') > \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n)" using assms cnf2bhv_def by auto text "Note that the functions are nat, that means that also in the case the difference is negative they will return a 0!" lemma cnf2bhv_ge_llength[simp]: assumes "n\\c \ t\" shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n) \ the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1" using assms cnf2bhv_def by simp lemma cnf2bhv_greater_llength[simp]: assumes "n>\c \ t\" shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n) > the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1" using assms cnf2bhv_def by simp lemma cnf2bhv_suc[simp]: assumes "n\\c \ t\" shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(Suc n) = Suc (\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n))" using assms cnf2bhv_def by simp lemma cnf2bhv_lActive[simp]: shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(\c \ t\) = the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1" using cnf2bhv_def by simp lemma cnf2bhv_lnth_lappend: assumes act: "\i. \c\\<^bsub>t i\<^esub>" and nAct: "\i. i\n \ \c\\<^bsub>t i\<^esub>" shows "lnth ((\\<^bsub>c\<^esub>(inf_llist t)) @\<^sub>l (inf_llist t')) (\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n)) = lnth (inf_llist t') (n - \c \ t\ - 1)" (is "?lhs = ?rhs") proof - from nAct have "lfinite (\\<^bsub>c\<^esub>(inf_llist t))" using proj_finite2 by auto then obtain k where k_def: "llength (\\<^bsub>c\<^esub>(inf_llist t)) = enat k" using lfinite_llength_enat by blast moreover have "k \ \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n)" proof - from nAct have "\i. i>n-1 \ \c\\<^bsub>t i\<^esub>" by simp with act have "\c \ t\ \ n-1" using lActive_less by auto moreover have "n>0" using act nAct by auto ultimately have "\c \ t\ < n" by simp hence "the_enat (llength (\\<^bsub>c\<^esub>inf_llist t)) - 1 < \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n)" using cnf2bhv_greater_llength by simp with k_def show ?thesis by simp qed ultimately have "?lhs = lnth (inf_llist t') (\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n) - k)" using lnth_lappend2 by blast moreover have "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n) - k = n - \c \ t\ - 1" proof - from cnf2bhv_def have "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n) - k = the_enat (llength (\\<^bsub>c\<^esub>inf_llist t)) - 1 + (n - \c \ t\) - k" by simp also have "\ = the_enat (llength (\\<^bsub>c\<^esub>inf_llist t)) - 1 + (n - \c \ t\) - the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t)))" using k_def by simp also have "\ = the_enat (llength (\\<^bsub>c\<^esub>inf_llist t)) + (n - \c \ t\) - 1 - the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t)))" proof - have "\i. enat i < llength (inf_llist t) \ \c\\<^bsub>lnth (inf_llist t) i\<^esub>" by (simp add: act) hence "llength (\\<^bsub>c\<^esub>inf_llist t) \ 1" using proj_one by simp moreover from k_def have "llength (\\<^bsub>c\<^esub>inf_llist t) \ \" by simp ultimately have "the_enat (llength (\\<^bsub>c\<^esub>inf_llist t)) \ 1" by (simp add: k_def one_enat_def) thus ?thesis by simp qed also have "\ = the_enat (llength (\\<^bsub>c\<^esub>inf_llist t)) + (n - \c \ t\) - the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1" by simp also have "\ = n - \c \ t\ - 1" by simp finally show ?thesis . qed ultimately show ?thesis by simp qed lemma nAct_cnf2proj_Suc_dist: assumes "\i\n. \c\\<^bsub>t i\<^esub>" and "\(\i>\c \ t\\<^bsub>n\<^esub>. \c\\<^bsub>t i\<^esub>)" shows "Suc (the_enat \c #\<^bsub>enat n\<^esub>inf_llist t\)=\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(Suc \c \ t\\<^bsub>n\<^esub>)" proof - have "the_enat \c #\<^bsub>enat n\<^esub>inf_llist t\ = \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(\c \ t\\<^bsub>n\<^esub>)" (is "?LHS = ?RHS") proof - from assms have "?RHS = the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1" using nxtActive_lactive[of n c t] by simp also have "llength (\\<^bsub>c\<^esub>(inf_llist t)) = eSuc (\c #\<^bsub>\c \ t\\<^bsub>n\<^esub>\<^esub> inf_llist t\)" proof - from assms have "\ (\i'\ Suc (\c \ t\\<^bsub>n\<^esub>). \c\\<^bsub>t i'\<^esub>)" using nxtActive_no_active by simp hence "\c #\<^bsub>Suc (\c \ t\\<^bsub>n\<^esub>)\<^esub> inf_llist t\ = llength (\\<^bsub>c\<^esub>(inf_llist t))" using nAct_eq_proj[of "Suc (\c \ t\\<^bsub>n\<^esub>)" c "inf_llist t"] by simp moreover from assms(1) have "\c\\<^bsub>t (\c \ t\\<^bsub>n\<^esub>)\<^esub>" using nxtActI by blast hence "\c #\<^bsub>Suc (\c \ t\\<^bsub>n\<^esub>)\<^esub> inf_llist t\ = eSuc (\c #\<^bsub>\c \ t\\<^bsub>n\<^esub>\<^esub> inf_llist t\)" by simp ultimately show ?thesis by simp qed also have "the_enat(eSuc (\c #\<^bsub>\c \ t\\<^bsub>n\<^esub>\<^esub> inf_llist t\)) - 1 = (\c #\<^bsub>\c \ t\\<^bsub>n\<^esub>\<^esub> inf_llist t\)" proof - have "\c #\<^bsub>\c \ t\\<^bsub>n\<^esub>\<^esub> inf_llist t\ \ \" by simp hence "the_enat(eSuc (\c #\<^bsub>\c \ t\\<^bsub>n\<^esub>\<^esub> inf_llist t\)) = Suc(the_enat(\c #\<^bsub>\c \ t\\<^bsub>n\<^esub>\<^esub> inf_llist t\))" using the_enat_eSuc by simp thus ?thesis by simp qed also have "\ = ?LHS" proof - have "enat \c \ t\\<^bsub>n\<^esub> - 1 < llength (inf_llist t)" by (simp add: one_enat_def) moreover from assms(1) have "\c \ t\\<^bsub>n\<^esub>\n" and "\k. enat n \ enat k \ enat k < enat \c \ t\\<^bsub>n\<^esub> \ \c\\<^bsub>lnth (inf_llist t) k\<^esub>" using nxtActI by auto ultimately have "\c #\<^bsub>enat \c \ t\\<^bsub>n\<^esub>\<^esub>inf_llist t\ = \c #\<^bsub>enat n\<^esub>inf_llist t\" using nAct_not_active_same[of n "\c \ t\\<^bsub>n\<^esub>" "inf_llist t" c] by simp moreover have "\c #\<^bsub>enat n\<^esub>inf_llist t\\\" by simp ultimately show ?thesis by auto qed finally show ?thesis by fastforce qed moreover from assms have "\c \ t\\<^bsub>n\<^esub>=\c \ t\" using nxtActive_lactive by simp hence "Suc (\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(\c \ t\\<^bsub>n\<^esub>)) = \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(Suc \c \ t\\<^bsub>n\<^esub>)" using cnf2bhv_suc[where n="\c \ t\\<^bsub>n\<^esub>"] by simp ultimately show ?thesis by simp qed subsubsection "Behavior Trace to Configuration Trace" text \ Next we define an operator to map a point in time of a behavior trace back to a corresponding point in time for a configuration trace. \ definition bhv2cnf :: "'id \ (nat \ cnf) \ nat \ nat" ("\<^bsub>_\<^esub>\\<^bsub>_\<^esub>(_)" [150,150,150] 110) where "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n) \ \c \ t\ + (n - (the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1))" lemma bhv2cnf_mono: assumes "n'\n" shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n') \ \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n)" by (simp add: assms bhv2cnf_def diff_le_mono) lemma bhv2cnf_mono_strict: assumes "n'>n" and "n \ the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1" shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n') > \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n)" using assms bhv2cnf_def by auto text "Note that the functions are nat, that means that also in the case the difference is negative they will return a 0!" lemma bhv2cnf_ge_lActive[simp]: shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n) \ \c \ t\" using bhv2cnf_def by simp lemma bhv2cnf_greater_lActive[simp]: assumes "n>the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1" shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n) > \c \ t\" using assms bhv2cnf_def by simp lemma bhv2cnf_lActive[simp]: assumes "\i. \c\\<^bsub>t i\<^esub>" and "lfinite (\\<^bsub>c\<^esub>(inf_llist t))" shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t)))) = Suc (\c \ t\)" proof - from assms have "\\<^bsub>c\<^esub>(inf_llist t)\ []\<^sub>l" by simp hence "llength (\\<^bsub>c\<^esub>(inf_llist t)) > 0" by (simp add: lnull_def) moreover from \lfinite (\\<^bsub>c\<^esub>(inf_llist t))\ have "llength (\\<^bsub>c\<^esub>(inf_llist t)) \ \" using llength_eq_infty_conv_lfinite by auto ultimately have "the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) > 0" using enat_0_iff(1) by fastforce hence "the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - (the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1) = 1" by simp thus ?thesis using bhv2cnf_def by simp qed subsubsection "Relating the Mappings" text \ In the following we provide some properties about the relationship between the two mapping operators. \ lemma bhv2cnf_cnf2bhv: assumes "n \ \c \ t\" shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n)) = n" (is "?lhs = ?rhs") proof - have "?lhs = \c \ t\ + ((\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n)) - (the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1))" using bhv2cnf_def by simp also have "\ = \c \ t\ + (((the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t)))) - 1 + (n - \c \ t\)) - (the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1))" using cnf2bhv_def by simp also have "(the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t)))) - 1 + (n - (\c \ t\)) - (the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1) = (the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t)))) - 1 - ((the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t)))) - 1) + (n - (\c \ t\))" by simp also have "\ = n - (\c \ t\)" by simp also have "(\c \ t\) + (n - (\c \ t\)) = (\c \ t\) + n - \c \ t\" using assms by simp also have "\ = ?rhs" by simp finally show ?thesis . qed lemma cnf2bhv_bhv2cnf: assumes "n \ the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1" shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n)) = n" (is "?lhs = ?rhs") proof - have "?lhs = the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1 + ((\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n)) - (\c \ t\))" using cnf2bhv_def by simp also have "\ = the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1 + (\c \ t\ + (n - (the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1)) - (\c \ t\))" using bhv2cnf_def by simp also have "\c \ t\ + (n - (the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1)) - (\c \ t\) = \c \ t\ - (\c \ t\) + (n - (the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1))" by simp also have "\ = n - (the_enat(llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1)" by simp also have "the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1 + (n - (the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1)) = n - (the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1) + (the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1)" by simp also have "\ = n + ((the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1) - (the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1))" using assms by simp also have "\ = ?rhs" by simp finally show ?thesis . qed lemma p2c_mono_c2p: assumes "n \ \c \ t\" and "n' \ \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n)" shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n') \ n" proof - from \n' \ \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n)\ have "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n') \ \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n))" using bhv2cnf_mono by simp thus ?thesis using bhv2cnf_cnf2bhv \n \ \c \ t\\ by simp qed lemma p2c_mono_c2p_strict: assumes "n \ \c \ t\" and "n<\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n')" shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n) < n'" proof (rule ccontr) assume "\ (\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n) < n')" hence "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n) \ n'" by simp with \n \ \c \ t\\ have "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(nat (\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n))) \ \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n')" using bhv2cnf_mono by simp hence "\(\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(nat (\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n))) < \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n'))" by simp with \n \ \c \ t\\ have "\(n < \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n'))" using "bhv2cnf_cnf2bhv" by simp with assms show False by simp qed lemma c2p_mono_p2c: assumes "n \ the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1" and "n' \ \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n)" shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n') \ n" proof - from \n' \ \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n)\ have "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n') \ \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n))" using cnf2bhv_mono by simp thus ?thesis using cnf2bhv_bhv2cnf \n \ the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1\ by simp qed lemma c2p_mono_p2c_strict: assumes "n \ the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1" and "n<\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n')" shows "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n) < n'" proof (rule ccontr) assume "\ (\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n) < n')" hence "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n) \ n'" by simp with \n \ the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1\ have "\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(nat (\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n))) \ \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n')" using cnf2bhv_mono by simp hence "\(\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(nat (\<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n))) < \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n'))" by simp with \n \ the_enat (llength (\\<^bsub>c\<^esub>(inf_llist t))) - 1\ have "\(n < \<^bsub>c\<^esub>\\<^bsub>t\<^esub>(n'))" using "cnf2bhv_bhv2cnf" by simp with assms show False by simp qed end end diff --git a/thys/QHLProver/Complex_Matrix.thy b/thys/QHLProver/Complex_Matrix.thy --- a/thys/QHLProver/Complex_Matrix.thy +++ b/thys/QHLProver/Complex_Matrix.thy @@ -1,2351 +1,2352 @@ section \Complex matrices\ theory Complex_Matrix imports "Jordan_Normal_Form.Matrix" "Jordan_Normal_Form.Conjugate" "Jordan_Normal_Form.Jordan_Normal_Form_Existence" begin subsection \Trace of a matrix\ definition trace :: "'a::ring mat \ 'a" where "trace A = (\ i \ {0 ..< dim_row A}. A $$ (i,i))" lemma trace_zero [simp]: "trace (0\<^sub>m n n) = 0" by (simp add: trace_def) lemma trace_id [simp]: "trace (1\<^sub>m n) = n" by (simp add: trace_def) lemma trace_comm: fixes A B :: "'a::comm_ring mat" assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" shows "trace (A * B) = trace (B * A)" proof (simp add: trace_def) have "(\i = 0..i = 0..j = 0.. = (\j = 0..i = 0.. = (\j = 0.. row B j)" by (metis (no_types, lifting) A B atLeastLessThan_iff carrier_matD index_col index_row scalar_prod_def sum.cong) also have "\ = (\j = 0.. col A j)" apply (rule sum.cong) apply auto apply (subst comm_scalar_prod[where n=n]) apply auto using assms by auto also have "\ = (\j = 0..i = 0..i = 0.. carrier_mat n n" and B: "B \ carrier_mat n n" shows "trace (A + B) = trace A + trace B" (is "?lhs = ?rhs") proof - have "?lhs = (\i=0.. = (\i=0..i=0..i=0..i=0..i=0..i=0.. carrier_mat n n" and B: "B \ carrier_mat n n" shows "trace (A - B) = trace A - trace B" (is "?lhs = ?rhs") proof - have "?lhs = (\i=0.. = (\i=0..i=0..i=0..i=0..i=0..i=0.. carrier_mat n n" shows "trace (c \\<^sub>m A) = c * trace A" proof - have "trace (c \\<^sub>m A) = (\i = 0.. = c * (\i = 0.. = c * trace A" unfolding trace_def by auto ultimately show ?thesis by auto qed subsection \Conjugate of a vector\ lemma conjugate_scalar_prod: fixes v w :: "'a::conjugatable_ring vec" assumes "dim_vec v = dim_vec w" shows "conjugate (v \ w) = conjugate v \ conjugate w" using assms by (simp add: scalar_prod_def sum_conjugate conjugate_dist_mul) subsection \Inner product\ abbreviation inner_prod :: "'a vec \ 'a vec \ 'a :: conjugatable_ring" where "inner_prod v w \ w \c v" lemma conjugate_scalar_prod_Im [simp]: "Im (v \c v) = 0" by (simp add: scalar_prod_def conjugate_vec_def sum.neutral) lemma conjugate_scalar_prod_Re [simp]: "Re (v \c v) \ 0" by (simp add: scalar_prod_def conjugate_vec_def sum_nonneg) lemma self_cscalar_prod_geq_0: fixes v :: "'a::conjugatable_ordered_field vec" shows "v \c v \ 0" by (auto simp add: scalar_prod_def, rule sum_nonneg, rule conjugate_square_positive) lemma inner_prod_distrib_left: fixes u v w :: "('a::conjugatable_field) vec" assumes dimu: "u \ carrier_vec n" and dimv:"v \ carrier_vec n" and dimw: "w \ carrier_vec n" shows "inner_prod (v + w) u = inner_prod v u + inner_prod w u" (is "?lhs = ?rhs") proof - have dimcv: "conjugate v \ carrier_vec n" and dimcw: "conjugate w \ carrier_vec n" using assms by auto have dimvw: "conjugate (v + w) \ carrier_vec n" using assms by auto have "u \ (conjugate (v + w)) = u \ conjugate v + u \ conjugate w" using dimv dimw dimu dimcv dimcw by (metis conjugate_add_vec scalar_prod_add_distrib) then show ?thesis by auto qed lemma inner_prod_distrib_right: fixes u v w :: "('a::conjugatable_field) vec" assumes dimu: "u \ carrier_vec n" and dimv:"v \ carrier_vec n" and dimw: "w \ carrier_vec n" shows "inner_prod u (v + w) = inner_prod u v + inner_prod u w" (is "?lhs = ?rhs") proof - have dimvw: "v + w \ carrier_vec n" using assms by auto have dimcu: "conjugate u \ carrier_vec n" using assms by auto have "(v + w) \ (conjugate u) = v \ conjugate u + w \ conjugate u" apply (simp add: comm_scalar_prod[OF dimvw dimcu]) apply (simp add: scalar_prod_add_distrib[OF dimcu dimv dimw]) apply (insert dimv dimw dimcu, simp add: comm_scalar_prod[of _ n]) done then show ?thesis by auto qed lemma inner_prod_minus_distrib_right: fixes u v w :: "('a::conjugatable_field) vec" assumes dimu: "u \ carrier_vec n" and dimv:"v \ carrier_vec n" and dimw: "w \ carrier_vec n" shows "inner_prod u (v - w) = inner_prod u v - inner_prod u w" (is "?lhs = ?rhs") proof - have dimvw: "v - w \ carrier_vec n" using assms by auto have dimcu: "conjugate u \ carrier_vec n" using assms by auto have "(v - w) \ (conjugate u) = v \ conjugate u - w \ conjugate u" apply (simp add: comm_scalar_prod[OF dimvw dimcu]) apply (simp add: scalar_prod_minus_distrib[OF dimcu dimv dimw]) apply (insert dimv dimw dimcu, simp add: comm_scalar_prod[of _ n]) done then show ?thesis by auto qed lemma inner_prod_smult_right: fixes u v :: "complex vec" assumes dimu: "u \ carrier_vec n" and dimv:"v \ carrier_vec n" shows "inner_prod (a \\<^sub>v u) v = conjugate a * inner_prod u v" (is "?lhs = ?rhs") using assms apply (simp add: scalar_prod_def conjugate_dist_mul) apply (subst sum_distrib_left) by (rule sum.cong, auto) lemma inner_prod_smult_left: fixes u v :: "complex vec" assumes dimu: "u \ carrier_vec n" and dimv: "v \ carrier_vec n" shows "inner_prod u (a \\<^sub>v v) = a * inner_prod u v" (is "?lhs = ?rhs") using assms apply (simp add: scalar_prod_def) apply (subst sum_distrib_left) by (rule sum.cong, auto) lemma inner_prod_smult_left_right: fixes u v :: "complex vec" assumes dimu: "u \ carrier_vec n" and dimv: "v \ carrier_vec n" shows "inner_prod (a \\<^sub>v u) (b \\<^sub>v v) = conjugate a * b * inner_prod u v" (is "?lhs = ?rhs") using assms apply (simp add: scalar_prod_def) apply (subst sum_distrib_left) by (rule sum.cong, auto) lemma inner_prod_swap: fixes x y :: "complex vec" assumes "y \ carrier_vec n" and "x \ carrier_vec n" shows "inner_prod y x = conjugate (inner_prod x y)" apply (simp add: scalar_prod_def) apply (rule sum.cong) using assms by auto text \Cauchy-Schwarz theorem for complex vectors. This is analogous to aux\_Cauchy and Cauchy\_Schwarz\_ineq in Generalizations2.thy in QR\_Decomposition. Consider merging and moving to Isabelle library.\ lemma aux_Cauchy: fixes x y :: "complex vec" assumes "x \ carrier_vec n" and "y \ carrier_vec n" shows "0 \ inner_prod x x + a * (inner_prod x y) + (cnj a) * ((cnj (inner_prod x y)) + a * (inner_prod y y))" proof - have "(inner_prod (x+ a \\<^sub>v y) (x+a \\<^sub>v y)) = (inner_prod (x+a \\<^sub>v y) x) + (inner_prod (x+a \\<^sub>v y) (a \\<^sub>v y))" apply (subst inner_prod_distrib_right) using assms by auto also have "\ = inner_prod x x + (a) * (inner_prod x y) + cnj a * ((cnj (inner_prod x y)) + (a) * (inner_prod y y))" apply (subst (1 2) inner_prod_distrib_left[of _ n]) apply (auto simp add: assms) apply (subst (1 2) inner_prod_smult_right[of _ n]) apply (auto simp add: assms) apply (subst inner_prod_smult_left[of _ n]) apply (auto simp add: assms) apply (subst inner_prod_swap[of y n x]) apply (auto simp add: assms) unfolding distrib_left by auto finally show ?thesis by (metis self_cscalar_prod_geq_0) qed lemma Cauchy_Schwarz_complex_vec: fixes x y :: "complex vec" assumes "x \ carrier_vec n" and "y \ carrier_vec n" shows "inner_prod x y * inner_prod y x \ inner_prod x x * inner_prod y y" proof - define cnj_a where "cnj_a = - (inner_prod x y)/ cnj (inner_prod y y)" define a where "a = cnj (cnj_a)" have cnj_rw: "(cnj a) = cnj_a" unfolding a_def by (simp) have rw_0: "cnj (inner_prod x y) + a * (inner_prod y y) = 0" unfolding a_def cnj_a_def using assms(1) assms(2) conjugate_square_eq_0_vec by fastforce have "0 \ (inner_prod x x + a * (inner_prod x y) + (cnj a) * ((cnj (inner_prod x y)) + a * (inner_prod y y)))" using aux_Cauchy assms by auto also have "\ = (inner_prod x x + a * (inner_prod x y))" unfolding rw_0 by auto also have "\ = (inner_prod x x - (inner_prod x y) * cnj (inner_prod x y) / (inner_prod y y))" unfolding a_def cnj_a_def by simp finally have " 0 \ (inner_prod x x - (inner_prod x y) * cnj (inner_prod x y) / (inner_prod y y)) " . hence "0 \ (inner_prod x x - (inner_prod x y) * cnj (inner_prod x y) / (inner_prod y y)) * (inner_prod y y)" by (auto simp: less_eq_complex_def) also have "\ = ((inner_prod x x)*(inner_prod y y) - (inner_prod x y) * cnj (inner_prod x y))" by (smt add.inverse_neutral add_diff_cancel diff_0 diff_divide_eq_iff divide_cancel_right mult_eq_0_iff nonzero_mult_div_cancel_right rw_0) finally have "(inner_prod x y) * cnj (inner_prod x y) \ (inner_prod x x)*(inner_prod y y)" by auto then show ?thesis apply (subst inner_prod_swap[of y n x]) by (auto simp add: assms) qed subsection \Hermitian adjoint of a matrix\ abbreviation adjoint where "adjoint \ mat_adjoint" lemma adjoint_dim_row [simp]: "dim_row (adjoint A) = dim_col A" by (simp add: mat_adjoint_def) lemma adjoint_dim_col [simp]: "dim_col (adjoint A) = dim_row A" by (simp add: mat_adjoint_def) lemma adjoint_dim: "A \ carrier_mat n n \ adjoint A \ carrier_mat n n" using adjoint_dim_col adjoint_dim_row by blast lemma adjoint_def: "adjoint A = mat (dim_col A) (dim_row A) (\(i,j). conjugate (A $$ (j,i)))" unfolding mat_adjoint_def mat_of_rows_def by auto lemma adjoint_eval: assumes "i < dim_col A" "j < dim_row A" shows "(adjoint A) $$ (i,j) = conjugate (A $$ (j,i))" using assms by (simp add: adjoint_def) lemma adjoint_row: assumes "i < dim_col A" shows "row (adjoint A) i = conjugate (col A i)" apply (rule eq_vecI) using assms by (auto simp add: adjoint_eval) lemma adjoint_col: assumes "i < dim_row A" shows "col (adjoint A) i = conjugate (row A i)" apply (rule eq_vecI) using assms by (auto simp add: adjoint_eval) text \The identity = \ lemma adjoint_def_alter: fixes v w :: "'a::conjugatable_field vec" and A :: "'a::conjugatable_field mat" assumes dims: "v \ carrier_vec n" "w \ carrier_vec m" "A \ carrier_mat n m" shows "inner_prod v (A *\<^sub>v w) = inner_prod (adjoint A *\<^sub>v v) w" (is "?lhs = ?rhs") proof - from dims have "?lhs = (\i=0..j=0..i=0..j=0..m n) = (1\<^sub>m n::complex mat)" apply (rule eq_matI) by (auto simp add: adjoint_eval) lemma adjoint_scale: fixes A :: "'a::conjugatable_field mat" shows "adjoint (a \\<^sub>m A) = (conjugate a) \\<^sub>m adjoint A" apply (rule eq_matI) using conjugatable_ring_class.conjugate_dist_mul by (auto simp add: adjoint_eval) lemma adjoint_add: fixes A B :: "'a::conjugatable_field mat" assumes "A \ carrier_mat n m" "B \ carrier_mat n m" shows "adjoint (A + B) = adjoint A + adjoint B" apply (rule eq_matI) using assms conjugatable_ring_class.conjugate_dist_add by( auto simp add: adjoint_eval) lemma adjoint_minus: fixes A B :: "'a::conjugatable_field mat" assumes "A \ carrier_mat n m" "B \ carrier_mat n m" shows "adjoint (A - B) = adjoint A - adjoint B" apply (rule eq_matI) using assms apply(auto simp add: adjoint_eval) by (metis add_uminus_conv_diff conjugate_dist_add conjugate_neg) lemma adjoint_mult: fixes A B :: "'a::conjugatable_field mat" assumes "A \ carrier_mat n m" "B \ carrier_mat m l" shows "adjoint (A * B) = adjoint B * adjoint A" proof (rule eq_matI, auto simp add: adjoint_eval adjoint_row adjoint_col) fix i j assume "i < dim_col B" "j < dim_row A" show "conjugate (row A j \ col B i) = conjugate (col B i) \ conjugate (row A j)" using assms apply (simp add: conjugate_scalar_prod) apply (subst comm_scalar_prod[where n="dim_row B"]) by (auto simp add: carrier_vecI) qed lemma adjoint_adjoint: fixes A :: "'a::conjugatable_field mat" shows "adjoint (adjoint A) = A" by (rule eq_matI, auto simp add: adjoint_eval) lemma trace_adjoint_positive: fixes A :: "complex mat" shows "trace (A * adjoint A) \ 0" apply (auto simp add: trace_def adjoint_col) apply (rule sum_nonneg) by auto subsection \Algebraic manipulations on matrices\ lemma right_add_zero_mat[simp]: "(A :: 'a :: monoid_add mat) \ carrier_mat nr nc \ A + 0\<^sub>m nr nc = A" by (intro eq_matI, auto) lemma add_carrier_mat': "A \ carrier_mat nr nc \ B \ carrier_mat nr nc \ A + B \ carrier_mat nr nc" by simp lemma minus_carrier_mat': "A \ carrier_mat nr nc \ B \ carrier_mat nr nc \ A - B \ carrier_mat nr nc" by auto lemma swap_plus_mat: fixes A B C :: "'a::semiring_1 mat" assumes "A \ carrier_mat n n" "B \ carrier_mat n n" "C \ carrier_mat n n" shows "A + B + C = A + C + B" by (metis assms assoc_add_mat comm_add_mat) lemma uminus_mat: fixes A :: "complex mat" assumes "A \ carrier_mat n n" shows "-A = (-1) \\<^sub>m A" by auto ML_file "mat_alg.ML" method_setup mat_assoc = \mat_assoc_method\ "Normalization of expressions on matrices" lemma mat_assoc_test: fixes A B C D :: "complex mat" assumes "A \ carrier_mat n n" "B \ carrier_mat n n" "C \ carrier_mat n n" "D \ carrier_mat n n" shows "(A * B) * (C * D) = A * B * C * D" "adjoint (A * adjoint B) * C = B * (adjoint A * C)" "A * 1\<^sub>m n * 1\<^sub>m n * B * 1\<^sub>m n = A * B" "(A - B) + (B - C) = A + (-B) + B + (-C)" "A + (B - C) = A + B - C" "A - (B + C + D) = A - B - C - D" "(A + B) * (B + C) = A * B + B * B + A * C + B * C" "A - B = A + (-1) \\<^sub>m B" "A * (B - C) * D = A * B * D - A * C * D" "trace (A * B * C) = trace (B * C * A)" "trace (A * B * C * D) = trace (C * D * A * B)" "trace (A + B * C) = trace A + trace (C * B)" "A + B = B + A" "A + B + C = C + B + A" "A + B + (C + D) = A + C + (B + D)" using assms by (mat_assoc n)+ subsection \Hermitian matrices\ text \A Hermitian matrix is a matrix that is equal to its Hermitian adjoint.\ definition hermitian :: "'a::conjugatable_field mat \ bool" where "hermitian A \ (adjoint A = A)" lemma hermitian_one: shows "hermitian ((1\<^sub>m n)::('a::conjugatable_field mat))" unfolding hermitian_def proof- have "conjugate (1::'a) = 1" apply (subst mult_1_right[symmetric, of "conjugate 1"]) apply (subst conjugate_id[symmetric, of "conjugate 1 * 1"]) apply (subst conjugate_dist_mul) apply auto done then show "adjoint ((1\<^sub>m n)::('a::conjugatable_field mat)) = (1\<^sub>m n)" by (auto simp add: adjoint_eval) qed subsection \Inverse matrices\ lemma inverts_mat_symm: fixes A B :: "'a::field mat" assumes dim: "A \ carrier_mat n n" "B \ carrier_mat n n" and AB: "inverts_mat A B" shows "inverts_mat B A" proof - have "A * B = 1\<^sub>m n" using dim AB unfolding inverts_mat_def by auto with dim have "B * A = 1\<^sub>m n" by (rule mat_mult_left_right_inverse) then show "inverts_mat B A" using dim inverts_mat_def by auto qed lemma inverts_mat_unique: fixes A B C :: "'a::field mat" assumes dim: "A \ carrier_mat n n" "B \ carrier_mat n n" "C \ carrier_mat n n" and AB: "inverts_mat A B" and AC: "inverts_mat A C" shows "B = C" proof - have AB1: "A * B = 1\<^sub>m n" using AB dim unfolding inverts_mat_def by auto have "A * C = 1\<^sub>m n" using AC dim unfolding inverts_mat_def by auto then have CA1: "C * A = 1\<^sub>m n" using mat_mult_left_right_inverse[of A n C] dim by auto then have "C = C * 1\<^sub>m n" using dim by auto also have "\ = C * (A * B)" using AB1 by auto also have "\ = (C * A) * B" using dim by auto also have "\ = 1\<^sub>m n * B" using CA1 by auto also have "\ = B" using dim by auto finally show "B = C" .. qed subsection \Unitary matrices\ text \A unitary matrix is a matrix whose Hermitian adjoint is also its inverse.\ definition unitary :: "'a::conjugatable_field mat \ bool" where "unitary A \ A \ carrier_mat (dim_row A) (dim_row A) \ inverts_mat A (adjoint A)" lemma unitaryD2: assumes "A \ carrier_mat n n" shows "unitary A \ inverts_mat (adjoint A) A" using assms adjoint_dim inverts_mat_symm unitary_def by blast lemma unitary_simps [simp]: "A \ carrier_mat n n \ unitary A \ adjoint A * A = 1\<^sub>m n" "A \ carrier_mat n n \ unitary A \ A * adjoint A = 1\<^sub>m n" apply (metis adjoint_dim_row carrier_matD(2) inverts_mat_def unitaryD2) by (simp add: inverts_mat_def unitary_def) lemma unitary_adjoint [simp]: assumes "A \ carrier_mat n n" "unitary A" shows "unitary (adjoint A)" unfolding unitary_def using adjoint_dim[OF assms(1)] assms by (auto simp add: unitaryD2[OF assms] adjoint_adjoint) lemma unitary_one: shows "unitary ((1\<^sub>m n)::('a::conjugatable_field mat))" unfolding unitary_def proof - define I where I_def[simp]: "I \ ((1\<^sub>m n)::('a::conjugatable_field mat))" have dim: "I \ carrier_mat n n" by auto have "hermitian I" using hermitian_one by auto hence "adjoint I = I" using hermitian_def by auto with dim show "I \ carrier_mat (dim_row I) (dim_row I) \ inverts_mat I (adjoint I)" unfolding inverts_mat_def using dim by auto qed lemma unitary_zero: fixes A :: "'a::conjugatable_field mat" assumes "A \ carrier_mat 0 0" shows "unitary A" unfolding unitary_def inverts_mat_def Let_def using assms by auto lemma unitary_elim: assumes dims: "A \ carrier_mat n n" "B \ carrier_mat n n" "P \ carrier_mat n n" and uP: "unitary P" and eq: "P * A * adjoint P = P * B * adjoint P" shows "A = B" proof - have dimaP: "adjoint P \ carrier_mat n n" using dims by auto have iv: "inverts_mat P (adjoint P)" using uP unitary_def by auto then have "P * (adjoint P) = 1\<^sub>m n" using inverts_mat_def dims by auto then have aPP: "adjoint P * P = 1\<^sub>m n" using mat_mult_left_right_inverse[OF dims(3) dimaP] by auto have "adjoint P * (P * A * adjoint P) * P = (adjoint P * P) * A * (adjoint P * P)" using dims dimaP by (mat_assoc n) also have "\ = 1\<^sub>m n * A * 1\<^sub>m n" using aPP by auto also have "\ = A" using dims by auto finally have eqA: "A = adjoint P * (P * A * adjoint P) * P" .. have "adjoint P * (P * B * adjoint P) * P = (adjoint P * P) * B * (adjoint P * P)" using dims dimaP by (mat_assoc n) also have "\ = 1\<^sub>m n * B * 1\<^sub>m n" using aPP by auto also have "\ = B" using dims by auto finally have eqB: "B = adjoint P * (P * B * adjoint P) * P" .. then show ?thesis using eqA eqB eq by auto qed lemma unitary_is_corthogonal: fixes U :: "'a::conjugatable_field mat" assumes dim: "U \ carrier_mat n n" and U: "unitary U" shows "corthogonal_mat U" unfolding corthogonal_mat_def Let_def proof (rule conjI) have dima: "adjoint U \ carrier_mat n n" using dim by auto have aUU: "mat_adjoint U * U = (1\<^sub>m n)" apply (insert U[unfolded unitary_def] dim dima, drule conjunct2) apply (drule inverts_mat_symm[of "U", OF dim dima], unfold inverts_mat_def, auto) done then show "diagonal_mat (mat_adjoint U * U)" by (simp add: diagonal_mat_def) show "\i 0" using dim by (simp add: aUU) qed lemma unitary_times_unitary: fixes P Q :: "'a:: conjugatable_field mat" assumes dim: "P \ carrier_mat n n" "Q \ carrier_mat n n" and uP: "unitary P" and uQ: "unitary Q" shows "unitary (P * Q)" proof - have dim_pq: "P * Q \ carrier_mat n n" using dim by auto have "(P * Q) * adjoint (P * Q) = P * (Q * adjoint Q) * adjoint P" using dim by (mat_assoc n) also have "\ = P * (1\<^sub>m n) * adjoint P" using uQ dim by auto also have "\ = P * adjoint P" using dim by (mat_assoc n) also have "\ = 1\<^sub>m n" using uP dim by simp finally have "(P * Q) * adjoint (P * Q) = 1\<^sub>m n" by auto hence "inverts_mat (P * Q) (adjoint (P * Q))" using inverts_mat_def dim_pq by auto thus "unitary (P*Q)" using unitary_def dim_pq by auto qed lemma unitary_operator_keep_trace: fixes U A :: "complex mat" assumes dU: "U \ carrier_mat n n" and dA: "A \ carrier_mat n n" and u: "unitary U" shows "trace A = trace (adjoint U * A * U)" proof - have u': "U * adjoint U = 1\<^sub>m n" using u unfolding unitary_def inverts_mat_def using dU by auto have "trace (adjoint U * A * U) = trace (U * adjoint U * A)" using dU dA by (mat_assoc n) also have "\ = trace A" using u' dA by auto finally show ?thesis by auto qed subsection \Normalization of vectors\ definition vec_norm :: "complex vec \ complex" where "vec_norm v \ csqrt (v \c v)" lemma vec_norm_geq_0: fixes v :: "complex vec" shows "vec_norm v \ 0" unfolding vec_norm_def by (insert self_cscalar_prod_geq_0[of v], simp add: less_eq_complex_def) lemma vec_norm_zero: fixes v :: "complex vec" assumes dim: "v \ carrier_vec n" shows "vec_norm v = 0 \ v = 0\<^sub>v n" unfolding vec_norm_def by (subst conjugate_square_eq_0_vec[OF dim, symmetric], rule csqrt_eq_0) lemma vec_norm_ge_0: fixes v :: "complex vec" assumes dim_v: "v \ carrier_vec n" and neq0: "v \ 0\<^sub>v n" shows "vec_norm v > 0" proof - have geq: "vec_norm v \ 0" using vec_norm_geq_0 by auto have neq: "vec_norm v \ 0" apply (insert dim_v neq0) apply (drule vec_norm_zero, auto) done show ?thesis using neq geq by (rule dual_order.not_eq_order_implies_strict) qed definition vec_normalize :: "complex vec \ complex vec" where "vec_normalize v = (if (v = 0\<^sub>v (dim_vec v)) then v else 1 / (vec_norm v) \\<^sub>v v)" lemma normalized_vec_dim[simp]: assumes "(v::complex vec) \ carrier_vec n" shows "vec_normalize v \ carrier_vec n" unfolding vec_normalize_def using assms by auto lemma vec_eq_norm_smult_normalized: shows "v = vec_norm v \\<^sub>v vec_normalize v" proof (cases "v = 0\<^sub>v (dim_vec v)") define n where "n = dim_vec v" then have dimv: "v \ carrier_vec n" by auto then have dimnv: "vec_normalize v \ carrier_vec n" by auto { case True then have v0: "v = 0\<^sub>v n" using n_def by auto then have n0: "vec_norm v = 0" using vec_norm_def by auto have "vec_norm v \\<^sub>v vec_normalize v = 0\<^sub>v n" unfolding smult_vec_def by (auto simp add: n0 carrier_vecD[OF dimnv]) then show ?thesis using v0 by auto next case False then have v: "v \ 0\<^sub>v n" using n_def by auto then have ge0: "vec_norm v > 0" using vec_norm_ge_0 dimv by auto have "vec_normalize v = (1 / vec_norm v) \\<^sub>v v" using False vec_normalize_def by auto then have "vec_norm v \\<^sub>v vec_normalize v = (vec_norm v * (1 / vec_norm v)) \\<^sub>v v" using smult_smult_assoc by auto also have "\ = v" using ge0 by auto finally have "v = vec_norm v \\<^sub>v vec_normalize v".. then show "v = vec_norm v \\<^sub>v vec_normalize v" using v by auto } qed lemma normalized_cscalar_prod: fixes v w :: "complex vec" assumes dim_v: "v \ carrier_vec n" and dim_w: "w \ carrier_vec n" shows "v \c w = (vec_norm v * vec_norm w) * (vec_normalize v \c vec_normalize w)" unfolding vec_normalize_def apply (split if_split, split if_split) proof (intro conjI impI) note dim0 = dim_v dim_w have dim: "dim_vec v = n" "dim_vec w = n" using dim0 by auto { assume "w = 0\<^sub>v n" "v = 0\<^sub>v n" then have lhs: "v \c w = 0" by auto then moreover have rhs: "vec_norm v * vec_norm w * (v \c w) = 0" by auto ultimately have "v \c w = vec_norm v * vec_norm w * (v \c w)" by auto } with dim show "w = 0\<^sub>v (dim_vec w) \ v = 0\<^sub>v (dim_vec v) \ v \c w = vec_norm v * vec_norm w * (v \c w)" by auto { assume asm: "w = 0\<^sub>v n" "v \ 0\<^sub>v n" then have w0: "conjugate w = 0\<^sub>v n" by auto with dim0 have "(1 / vec_norm v \\<^sub>v v) \c w = 0" by auto then moreover have rhs: "vec_norm v * vec_norm w * ((1 / vec_norm v \\<^sub>v v) \c w) = 0" by auto moreover have "v \c w = 0" using w0 dim0 by auto ultimately have "v \c w = vec_norm v * vec_norm w * ((1 / vec_norm v \\<^sub>v v) \c w)" by auto } with dim show "w = 0\<^sub>v (dim_vec w) \ v \ 0\<^sub>v (dim_vec v) \ v \c w = vec_norm v * vec_norm w * ((1 / vec_norm v \\<^sub>v v) \c w)" by auto { assume asm: "w \ 0\<^sub>v n" "v = 0\<^sub>v n" with dim0 have "v \c (1 / vec_norm w \\<^sub>v w) = 0" by auto then moreover have rhs: "vec_norm v * vec_norm w * (v \c (1 / vec_norm w \\<^sub>v w)) = 0" by auto moreover have "v \c w = 0" using asm dim0 by auto ultimately have "v \c w = vec_norm v * vec_norm w * (v \c (1 / vec_norm w \\<^sub>v w))" by auto } with dim show "w \ 0\<^sub>v (dim_vec w) \ v = 0\<^sub>v (dim_vec v) \ v \c w = vec_norm v * vec_norm w * (v \c (1 / vec_norm w \\<^sub>v w))" by auto { assume asmw: "w \ 0\<^sub>v n" and asmv: "v \ 0\<^sub>v n" have "vec_norm w > 0" by (insert asmw dim0, rule vec_norm_ge_0, auto) then have cw: "conjugate (1 / vec_norm w) = 1 / vec_norm w" by (simp add: complex_eq_iff complex_is_Real_iff less_complex_def) from dim0 have "((1 / vec_norm v \\<^sub>v v) \c (1 / vec_norm w \\<^sub>v w)) = 1 / vec_norm v * (v \c (1 / vec_norm w \\<^sub>v w))" by auto also have "\ = 1 / vec_norm v * (v \ (conjugate (1 / vec_norm w) \\<^sub>v conjugate w))" by (subst conjugate_smult_vec, auto) also have "\ = 1 / vec_norm v * conjugate (1 / vec_norm w) * (v \ conjugate w)" using dim by auto also have "\ = 1 / vec_norm v * (1 / vec_norm w) * (v \c w)" using vec_norm_ge_0 cw by auto finally have eq1: "(1 / vec_norm v \\<^sub>v v) \c (1 / vec_norm w \\<^sub>v w) = 1 / vec_norm v * (1 / vec_norm w) * (v \c w)" . then have "vec_norm v * vec_norm w * ((1 / vec_norm v \\<^sub>v v) \c (1 / vec_norm w \\<^sub>v w)) = (v \c w)" by (subst eq1, insert vec_norm_ge_0[of v n, OF dim_v asmv] vec_norm_ge_0[of w n, OF dim_w asmw], auto) } with dim show " w \ 0\<^sub>v (dim_vec w) \ v \ 0\<^sub>v (dim_vec v) \ v \c w = vec_norm v * vec_norm w * ((1 / vec_norm v \\<^sub>v v) \c (1 / vec_norm w \\<^sub>v w))" by auto qed lemma normalized_vec_norm : fixes v :: "complex vec" assumes dim_v: "v \ carrier_vec n" and neq0: "v \ 0\<^sub>v n" shows "vec_normalize v \c vec_normalize v = 1" unfolding vec_normalize_def proof (simp, rule conjI) show "v = 0\<^sub>v (dim_vec v) \ v \c v = 1" using neq0 dim_v by auto have dim_a: "(vec_normalize v) \ carrier_vec n" "conjugate (vec_normalize v) \ carrier_vec n" using dim_v vec_normalize_def by auto note dim = dim_v dim_a have nvge0: "vec_norm v > 0" using vec_norm_ge_0 neq0 dim_v by auto then have vvvv: "v \c v = (vec_norm v) * (vec_norm v)" unfolding vec_norm_def by (metis power2_csqrt power2_eq_square) from nvge0 have "conjugate (vec_norm v) = vec_norm v" by (simp add: complex_eq_iff complex_is_Real_iff less_complex_def) then have "v \c (1 / vec_norm v \\<^sub>v v) = 1 / vec_norm v * (v \c v)" by (subst conjugate_smult_vec, auto) also have "\ = 1 / vec_norm v * vec_norm v * vec_norm v" using vvvv by auto also have "\ = vec_norm v" by auto finally have "v \c (1 / vec_norm v \\<^sub>v v) = vec_norm v". then show "v \ 0\<^sub>v (dim_vec v) \ vec_norm v \ 0 \ v \c (1 / vec_norm v \\<^sub>v v) = vec_norm v" using neq0 nvge0 by auto qed lemma normalize_zero: assumes "v \ carrier_vec n" shows "vec_normalize v = 0\<^sub>v n \ v = 0\<^sub>v n" proof show "v = 0\<^sub>v n \ vec_normalize v = 0\<^sub>v n" unfolding vec_normalize_def by auto next have "v \ 0\<^sub>v n \ vec_normalize v \ 0\<^sub>v n" unfolding vec_normalize_def proof (simp, rule impI) assume asm: "v \ 0\<^sub>v n" then have "vec_norm v > 0" using vec_norm_ge_0 assms by auto then have nvge0: "1 / vec_norm v > 0" by (simp add: complex_is_Real_iff less_complex_def) have "\k < n. v $ k \ 0" using asm assms by auto then obtain k where kn: "k < n" and vkneq0: "v $ k \ 0" by auto then have "(1 / vec_norm v \\<^sub>v v) $ k = (1 / vec_norm v) * (v $ k)" using assms carrier_vecD index_smult_vec(1) by blast with nvge0 vkneq0 have "(1 / vec_norm v \\<^sub>v v) $ k \ 0" by auto then show "1 / vec_norm v \\<^sub>v v \ 0\<^sub>v n" using assms kn by fastforce qed then show "vec_normalize v = 0\<^sub>v n \ v = 0\<^sub>v n" by auto qed lemma normalize_normalize[simp]: "vec_normalize (vec_normalize v) = vec_normalize v" proof (rule disjE[of "v = 0\<^sub>v (dim_vec v)" "v \ 0\<^sub>v (dim_vec v)"], auto) let ?n = "dim_vec v" { assume "v = 0\<^sub>v ?n" then have "vec_normalize v = v" unfolding vec_normalize_def by auto then show "vec_normalize (vec_normalize v) = vec_normalize v" by auto } assume neq0: "v \ 0\<^sub>v ?n" have dim: "v \ carrier_vec ?n" by auto have "vec_norm (vec_normalize v) = 1" unfolding vec_norm_def using normalized_vec_norm[OF dim neq0] by auto then show "vec_normalize (vec_normalize v) = vec_normalize v" by (subst (1) vec_normalize_def, simp) qed subsection \Spectral decomposition of normal complex matrices\ lemma normalize_keep_corthogonal: fixes vs :: "complex vec list" assumes cor: "corthogonal vs" and dims: "set vs \ carrier_vec n" shows "corthogonal (map vec_normalize vs)" unfolding corthogonal_def proof (rule allI, rule impI, rule allI, rule impI, goal_cases) case c: (1 i j) let ?m = "length vs" have len: "length (map vec_normalize vs) = ?m" by auto have dim: "\k. k < ?m \ (vs ! k) \ carrier_vec n" using dims by auto have map: "\k. k < ?m \ map vec_normalize vs ! k = vec_normalize (vs ! k)" by auto have eq1: "\j k. j < ?m \ k < ?m \ ((vs ! j) \c (vs ! k) = 0) = (j \ k)" using assms unfolding corthogonal_def by auto then have "\k. k < ?m \ (vs ! k) \c (vs ! k) \ 0 " by auto then have "\k. k < ?m \ (vs ! k) \ (0\<^sub>v n)" using dim by (auto simp add: conjugate_square_eq_0_vec[of _ n, OF dim]) then have vnneq0: "\k. k < ?m \ vec_norm (vs ! k) \ 0" using vec_norm_zero[OF dim] by auto then have i0: "vec_norm (vs ! i) \ 0" and j0: "vec_norm (vs ! j) \ 0" using c by auto have "(vs ! i) \c (vs ! j) = vec_norm (vs ! i) * vec_norm (vs ! j) * (vec_normalize (vs ! i) \c vec_normalize (vs ! j))" by (subst normalized_cscalar_prod[of "vs ! i" n "vs ! j"], auto, insert dim c, auto) with i0 j0 have "(vec_normalize (vs ! i) \c vec_normalize (vs ! j) = 0) = ((vs ! i) \c (vs ! j) = 0)" by auto with eq1 c have "(vec_normalize (vs ! i) \c vec_normalize (vs ! j) = 0) = (i \ j)" by auto with map c show "(map vec_normalize vs ! i \c map vec_normalize vs ! j = 0) = (i \ j)" by auto qed lemma normalized_corthogonal_mat_is_unitary: assumes W: "set ws \ carrier_vec n" and orth: "corthogonal ws" and len: "length ws = n" shows "unitary (mat_of_cols n (map vec_normalize ws))" (is "unitary ?W") proof - define vs where "vs = map vec_normalize ws" define W where "W = mat_of_cols n vs" have W': "set vs \ carrier_vec n" using assms vs_def by auto then have W'': "\k. k < length vs \ vs ! k \ carrier_vec n" by auto have orth': "corthogonal vs" using assms normalize_keep_corthogonal vs_def by auto have len'[simp]: "length vs = n" using assms vs_def by auto have dimW: "W \ carrier_mat n n" using W_def len by auto have "adjoint W \ carrier_mat n n" using dimW by auto then have dimaW: "mat_adjoint W \ carrier_mat n n" by auto { fix i j assume i: "i < n" and j: "j < n" have dimws: "(ws ! i) \ carrier_vec n" "(ws ! j) \ carrier_vec n" using W len i j by auto have "(ws ! i) \c (ws ! i) \ 0" "(ws ! j) \c (ws ! j) \ 0" using orth corthogonal_def[of ws] len i j by auto then have neq0: "(ws ! i) \ 0\<^sub>v n" "(ws ! j) \ 0\<^sub>v n" by (auto simp add: conjugate_square_eq_0_vec[of "ws ! i" n]) then have "vec_norm (ws ! i) > 0" "vec_norm (ws ! j) > 0" using vec_norm_ge_0 dimws by auto then have ge0: "vec_norm (ws ! i) * vec_norm (ws ! j) > 0" by (auto simp: less_complex_def) have ws': "vs ! i = vec_normalize (ws ! i)" "vs ! j = vec_normalize (ws ! j)" using len i j vs_def by auto have ii1: "(vs ! i) \c (vs ! i) = 1" apply (simp add: ws') apply (rule normalized_vec_norm[of "ws ! i"], rule dimws, rule neq0) done have ij0: "i \ j \ (ws ! i) \c (ws ! j) = 0" using i j by (insert orth, auto simp add: corthogonal_def[of ws] len) have "i \ j \ (ws ! i) \c (ws ! j) = (vec_norm (ws ! i) * vec_norm (ws ! j)) * ((vs ! i) \c (vs ! j))" apply (auto simp add: ws') apply (rule normalized_cscalar_prod) apply (rule dimws, rule dimws) done with ij0 have ij0': "i \ j \ (vs ! i) \c (vs ! j) = 0" using ge0 by auto have cWk: "\k. k < n \ col W k = vs ! k" unfolding W_def apply (subst col_mat_of_cols) apply (auto simp add: W'') done have "(mat_adjoint W * W) $$ (j, i) = row (mat_adjoint W) j \ col W i" by (insert dimW i j dimaW, auto) also have "\ = conjugate (col W j) \ col W i" by (insert dimW i j dimaW, auto simp add: mat_adjoint_def) also have "\ = col W i \ conjugate (col W j)" using comm_scalar_prod[of "col W i" n] dimW by auto also have "\ = (vs ! i) \c (vs ! j)" using W_def col_mat_of_cols i j len cWk by auto finally have "(mat_adjoint W * W) $$ (j, i) = (vs ! i) \c (vs ! j)". then have "(mat_adjoint W * W) $$ (j, i) = (if (j = i) then 1 else 0)" by (auto simp add: ii1 ij0') } note maWW = this then have "mat_adjoint W * W = 1\<^sub>m n" unfolding one_mat_def using dimW dimaW by (auto simp add: maWW adjoint_def) then have iv0: "adjoint W * W = 1\<^sub>m n" by auto have dimaW: "adjoint W \ carrier_mat n n" using dimaW by auto then have iv1: "W * adjoint W = 1\<^sub>m n" using mat_mult_left_right_inverse dimW iv0 by auto then show "unitary W" unfolding unitary_def inverts_mat_def using dimW dimaW iv0 iv1 by auto qed lemma normalize_keep_eigenvector: assumes ev: "eigenvector A v e" and dim: "A \ carrier_mat n n" "v \ carrier_vec n" shows "eigenvector A (vec_normalize v) e" unfolding eigenvector_def proof show "vec_normalize v \ carrier_vec (dim_row A)" using dim by auto have eg: "A *\<^sub>v v = e \\<^sub>v v" using ev dim eigenvector_def by auto have vneq0: "v \ 0\<^sub>v n" using ev dim unfolding eigenvector_def by auto then have s0: "vec_normalize v \ 0\<^sub>v n" by (insert dim, subst normalize_zero[of v], auto) from vneq0 have vvge0: "vec_norm v > 0" using vec_norm_ge_0 dim by auto have s1: "A *\<^sub>v vec_normalize v = e \\<^sub>v vec_normalize v" unfolding vec_normalize_def using vneq0 dim apply (auto, simp add: mult_mat_vec) apply (subst eg, auto) done with s0 dim show "vec_normalize v \ 0\<^sub>v (dim_row A) \ A *\<^sub>v vec_normalize v = e \\<^sub>v vec_normalize v" by auto qed lemma four_block_mat_adjoint: fixes A B C D :: "'a::conjugatable_field mat" assumes dim: "A \ carrier_mat nr1 nc1" "B \ carrier_mat nr1 nc2" "C \ carrier_mat nr2 nc1" "D \ carrier_mat nr2 nc2" shows "adjoint (four_block_mat A B C D) = four_block_mat (adjoint A) (adjoint C) (adjoint B) (adjoint D)" by (rule eq_matI, insert dim, auto simp add: adjoint_eval) fun unitary_schur_decomposition :: "complex mat \ complex list \ complex mat \ complex mat \ complex mat" where "unitary_schur_decomposition A [] = (A, 1\<^sub>m (dim_row A), 1\<^sub>m (dim_row A))" | "unitary_schur_decomposition A (e # es) = (let n = dim_row A; n1 = n - 1; v' = find_eigenvector A e; v = vec_normalize v'; ws0 = gram_schmidt n (basis_completion v); ws = map vec_normalize ws0; W = mat_of_cols n ws; W' = corthogonal_inv W; A' = W' * A * W; (A1,A2,A0,A3) = split_block A' 1 1; (B,P,Q) = unitary_schur_decomposition A3 es; z_row = (0\<^sub>m 1 n1); z_col = (0\<^sub>m n1 1); one_1 = 1\<^sub>m 1 in (four_block_mat A1 (A2 * P) A0 B, W * four_block_mat one_1 z_row z_col P, four_block_mat one_1 z_row z_col Q * W'))" theorem unitary_schur_decomposition: assumes A: "(A::complex mat) \ carrier_mat n n" and c: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" and B: "unitary_schur_decomposition A es = (B,P,Q)" shows "similar_mat_wit A B P Q \ upper_triangular B \ diag_mat B = es \ unitary P \ (Q = adjoint P)" using assms proof (induct es arbitrary: n A B P Q) case Nil with degree_monic_char_poly[of A n] show ?case by (auto intro: similar_mat_wit_refl simp: diag_mat_def unitary_zero) next case (Cons e es n A C P Q) let ?n1 = "n - 1" from Cons have A: "A \ carrier_mat n n" and dim: "dim_row A = n" by auto let ?cp = "char_poly A" from Cons(3) have cp: "?cp = [: -e, 1 :] * (\e \ es. [:- e, 1:])" by auto have mon: "monic (\e\ es. [:- e, 1:])" by (rule monic_prod_list, auto) have deg: "degree ?cp = Suc (degree (\e\ es. [:- e, 1:]))" unfolding cp by (subst degree_mult_eq, insert mon, auto) with degree_monic_char_poly[OF A] have n: "n \ 0" by auto define v' where "v' = find_eigenvector A e" define v where "v = vec_normalize v'" define b where "b = basis_completion v" define ws0 where "ws0 = gram_schmidt n b" define ws where "ws = map vec_normalize ws0" define W where "W = mat_of_cols n ws" define W' where "W' = corthogonal_inv W" define A' where "A' = W' * A * W" obtain A1 A2 A0 A3 where splitA': "split_block A' 1 1 = (A1,A2,A0,A3)" by (cases "split_block A' 1 1", auto) obtain B P' Q' where schur: "unitary_schur_decomposition A3 es = (B,P',Q')" by (cases "unitary_schur_decomposition A3 es", auto) let ?P' = "four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 ?n1) (0\<^sub>m ?n1 1) P'" let ?Q' = "four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 ?n1) (0\<^sub>m ?n1 1) Q'" have C: "C = four_block_mat A1 (A2 * P') A0 B" and P: "P = W * ?P'" and Q: "Q = ?Q' * W'" using Cons(4) unfolding unitary_schur_decomposition.simps Let_def list.sel dim v'_def[symmetric] v_def[symmetric] b_def[symmetric] ws0_def[symmetric] ws_def[symmetric] W'_def[symmetric] W_def[symmetric] A'_def[symmetric] split splitA' schur by auto have e: "eigenvalue A e" unfolding eigenvalue_root_char_poly[OF A] cp by simp from find_eigenvector[OF A e] have ev': "eigenvector A v' e" unfolding v'_def . then have "v' \ carrier_vec n" unfolding eigenvector_def using A by auto with ev' have ev: "eigenvector A v e" unfolding v_def using A dim normalize_keep_eigenvector by auto from this[unfolded eigenvector_def] have v[simp]: "v \ carrier_vec n" and v0: "v \ 0\<^sub>v n" using A by auto interpret cof_vec_space n "TYPE(complex)" . from basis_completion[OF v v0, folded b_def] have span_b: "span (set b) = carrier_vec n" and dist_b: "distinct b" and indep: "\ lin_dep (set b)" and b: "set b \ carrier_vec n" and hdb: "hd b = v" and len_b: "length b = n" by auto from hdb len_b n obtain vs where bv: "b = v # vs" by (cases b, auto) from gram_schmidt_result[OF b dist_b indep refl, folded ws0_def] have ws0: "set ws0 \ carrier_vec n" "corthogonal ws0" "length ws0 = n" by (auto simp: len_b) then have ws: "set ws \ carrier_vec n" "corthogonal ws" "length ws = n" unfolding ws_def using normalize_keep_corthogonal by auto have ws0ne: "ws0 \ []" using \length ws0 = n\ n by auto from gram_schmidt_hd[OF v, of vs, folded bv] have hdws0: "hd ws0 = (vec_normalize v')" unfolding ws0_def v_def . have "hd ws = vec_normalize (hd ws0)" unfolding ws_def using hd_map[OF ws0ne] by auto then have hdws: "hd ws = v" unfolding v_def using normalize_normalize[of v'] hdws0 by auto have orth_W: "corthogonal_mat W" using orthogonal_mat_of_cols ws unfolding W_def. have W: "W \ carrier_mat n n" using ws unfolding W_def using mat_of_cols_carrier(1)[of n ws] by auto have W': "W' \ carrier_mat n n" unfolding W'_def corthogonal_inv_def using W by (auto simp: mat_of_rows_def) from corthogonal_inv_result[OF orth_W] have W'W: "inverts_mat W' W" unfolding W'_def . hence WW': "inverts_mat W W'" using mat_mult_left_right_inverse[OF W' W] W' W unfolding inverts_mat_def by auto have A': "A' \ carrier_mat n n" using W W' A unfolding A'_def by auto have A'A_wit: "similar_mat_wit A' A W' W" by (rule similar_mat_witI[of _ _ n], insert W W' A A' W'W WW', auto simp: A'_def inverts_mat_def) hence A'A: "similar_mat A' A" unfolding similar_mat_def by blast from similar_mat_wit_sym[OF A'A_wit] have simAA': "similar_mat_wit A A' W W'" by auto have eigen[simp]: "A *\<^sub>v v = e \\<^sub>v v" and v0: "v \ 0\<^sub>v n" using v_def v'_def find_eigenvector[OF A e] A normalize_keep_eigenvector unfolding eigenvector_def by auto let ?f = "(\ i. if i = 0 then e else 0)" have col0: "col A' 0 = vec n ?f" unfolding A'_def W'_def W_def using corthogonal_col_ev_0[OF A v v0 eigen n hdws ws]. from A' n have "dim_row A' = 1 + ?n1" "dim_col A' = 1 + ?n1" by auto from split_block[OF splitA' this] have A2: "A2 \ carrier_mat 1 ?n1" and A3: "A3 \ carrier_mat ?n1 ?n1" and A'block: "A' = four_block_mat A1 A2 A0 A3" by auto have A1id: "A1 = mat 1 1 (\ _. e)" using splitA'[unfolded split_block_def Let_def] arg_cong[OF col0, of "\ v. v $ 0"] A' n by (auto simp: col_def) have A1: "A1 \ carrier_mat 1 1" unfolding A1id by auto { fix i assume "i < ?n1" with arg_cong[OF col0, of "\ v. v $ Suc i"] A' have "A' $$ (Suc i, 0) = 0" by auto } note A'0 = this have A0id: "A0 = 0\<^sub>m ?n1 1" using splitA'[unfolded split_block_def Let_def] A'0 A' by auto have A0: "A0 \ carrier_mat ?n1 1" unfolding A0id by auto from cp char_poly_similar[OF A'A] have cp: "char_poly A' = [: -e,1 :] * (\ e \ es. [:- e, 1:])" by simp also have "char_poly A' = char_poly A1 * char_poly A3" unfolding A'block A0id by (rule char_poly_four_block_zeros_col[OF A1 A2 A3]) also have "char_poly A1 = [: -e,1 :]" by (simp add: A1id char_poly_defs det_def) finally have cp: "char_poly A3 = (\ e \ es. [:- e, 1:])" by (metis mult_cancel_left pCons_eq_0_iff zero_neq_one) from Cons(1)[OF A3 cp schur] have simIH: "similar_mat_wit A3 B P' Q'" and ut: "upper_triangular B" and diag: "diag_mat B = es" and uP': "unitary P'" and Q'P': "Q' = adjoint P'" by auto from similar_mat_witD2[OF A3 simIH] have B: "B \ carrier_mat ?n1 ?n1" and P': "P' \ carrier_mat ?n1 ?n1" and Q': "Q' \ carrier_mat ?n1 ?n1" and PQ': "P' * Q' = 1\<^sub>m ?n1" by auto have A0_eq: "A0 = P' * A0 * 1\<^sub>m 1" unfolding A0id using P' by auto have simA'C: "similar_mat_wit A' C ?P' ?Q'" unfolding A'block C by (rule similar_mat_wit_four_block[OF similar_mat_wit_refl[OF A1] simIH _ A0_eq A1 A3 A0], insert PQ' A2 P' Q', auto) have ut1: "upper_triangular A1" unfolding A1id by auto have ut: "upper_triangular C" unfolding C A0id by (intro upper_triangular_four_block[OF _ B ut1 ut], auto simp: A1id) from A1id have diagA1: "diag_mat A1 = [e]" unfolding diag_mat_def by auto from diag_four_block_mat[OF A1 B] have diag: "diag_mat C = e # es" unfolding diag diagA1 C by simp have aW: "adjoint W \ carrier_mat n n" using W by auto have aW': "adjoint W' \ carrier_mat n n" using W' by auto have "unitary W" using W_def ws_def ws0 normalized_corthogonal_mat_is_unitary by auto then have ivWaW: "inverts_mat W (adjoint W)" using unitary_def W aW by auto with WW' have W'aW: "W' = (adjoint W)" using inverts_mat_unique W W' aW by auto then have "adjoint W' = W" using adjoint_adjoint by auto with ivWaW have "inverts_mat W' (adjoint W')" using inverts_mat_symm W aW W'aW by auto then have "unitary W'" using unitary_def W' by auto have newP': "P' \ carrier_mat (n - Suc 0) (n - Suc 0)" using P' by auto have rl: "\ x1 x2 x3 x4 y1 y2 y3 y4 f. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ f x1 x2 x3 x4 = f y1 y2 y3 y4" by simp have Q'aP': "?Q' = adjoint ?P'" apply (subst four_block_mat_adjoint, auto simp add: newP') apply (rule rl[where f2 = four_block_mat]) apply (auto simp add: eq_matI adjoint_eval Q'P') done have "adjoint P = adjoint ?P' * adjoint W" using W newP' n apply (simp add: P) apply (subst adjoint_mult[of W, symmetric]) apply (auto simp add: W P' carrier_matD[of W n n]) done also have "\ = ?Q' * W'" using Q'aP' W'aW by auto also have "\ = Q" using Q by auto finally have QaP: "Q = adjoint P" .. from similar_mat_wit_trans[OF simAA' simA'C, folded P Q] have smw: "similar_mat_wit A C P Q" by blast then have dimP: "P \ carrier_mat n n" and dimQ: "Q \ carrier_mat n n" unfolding similar_mat_wit_def using A by auto from smw have "P * Q = 1\<^sub>m n" unfolding similar_mat_wit_def using A by auto then have "inverts_mat P Q" using inverts_mat_def dimP by auto then have uP: "unitary P" using QaP unitary_def dimP by auto from ut similar_mat_wit_trans[OF simAA' simA'C, folded P Q] diag uP QaP show ?case by blast qed lemma complex_mat_char_poly_factorizable: fixes A :: "complex mat" assumes "A \ carrier_mat n n" shows "\as. char_poly A = (\ a \ as. [:- a, 1:]) \ length as = n" proof - let ?ca = "char_poly A" have ex0: "\bs. Polynomial.smult (lead_coeff ?ca) (\b\bs. [:- b, 1:]) = ?ca \ length bs = degree ?ca" by (simp add: fundamental_theorem_algebra_factorized) then obtain bs where " Polynomial.smult (lead_coeff ?ca) (\b\bs. [:- b, 1:]) = ?ca \ length bs = degree ?ca" by auto moreover have "lead_coeff ?ca = (1::complex)" using assms degree_monic_char_poly by blast ultimately have ex1: "?ca = (\b\bs. [:- b, 1:]) \ length bs = degree ?ca" by auto moreover have "degree ?ca = n" by (simp add: assms degree_monic_char_poly) ultimately show ?thesis by auto qed lemma complex_mat_has_unitary_schur_decomposition: fixes A :: "complex mat" assumes "A \ carrier_mat n n" shows "\B P es. similar_mat_wit A B P (adjoint P) \ unitary P \ char_poly A = (\ (e :: complex) \ es. [:- e, 1:]) \ diag_mat B = es" proof - have "\es. char_poly A = (\ e \ es. [:- e, 1:]) \ length es = n" using assms by (simp add: complex_mat_char_poly_factorizable) then obtain es where es: "char_poly A = (\ e \ es. [:- e, 1:]) \ length es = n" by auto obtain B P Q where B: "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) have "similar_mat_wit A B P Q \ upper_triangular B \ unitary P \ (Q = adjoint P) \ char_poly A = (\ (e :: complex) \ es. [:- e, 1:]) \ diag_mat B = es" using assms es B by (auto simp add: unitary_schur_decomposition) then show ?thesis by auto qed lemma normal_upper_triangular_matrix_is_diagonal: fixes A :: "'a::conjugatable_ordered_field mat" assumes "A \ carrier_mat n n" and tri: "upper_triangular A" and norm: "A * adjoint A = adjoint A * A" shows "diagonal_mat A" proof (rule disjE[of "n = 0" "n > 0"], blast) have dim: "dim_row A = n" "dim_col A = n" using assms by auto from norm have eq0: "\i j. (A * adjoint A)$$(i,j) = (adjoint A * A)$$(i,j)" by auto have nat_induct_strong: "\P. (P::nat\bool) 0 \ (\i. i < n \ (\k. k < i \ P k) \ P i) \ (\i. i < n \ P i)" by (metis dual_order.strict_trans infinite_descent0 linorder_neqE_nat) show "n = 0 \ ?thesis" using dim unfolding diagonal_mat_def by auto show "n > 0 \ ?thesis" unfolding diagonal_mat_def dim apply (rule allI, rule impI) apply (rule nat_induct_strong) proof (rule allI, rule impI, rule impI) assume asm: "n > 0" from tri upper_triangularD[of A 0 j] dim have z0: "\j. 0< j \ j < n \ A$$(j, 0) = 0" by auto then have ada00: "(adjoint A * A)$$(0,0) = conjugate (A$$(0,0)) * A$$(0,0)" using asm dim by (auto simp add: scalar_prod_def adjoint_eval sum.atLeast_Suc_lessThan) have aad00: "(A * adjoint A)$$(0,0) = (\k=0.. = A$$(0,0) * conjugate (A$$(0,0)) + (\k=1..k. A$$(0, k) * conjugate (A$$(0, k))"], auto) ultimately have f1tneq0: "(\k=(Suc 0)..k. k < n \ A$$(0, k) * conjugate (A$$(0, k)) \ 0" using conjugate_square_positive by auto have "\k. 1 \ k \ k < n \ A$$(0, k) * conjugate (A$$(0, k)) = 0" by (rule sum_nonneg_0[of "{1..j. 0 < n \ j < n \ 0 \ j \ A $$ (0, j) = 0" by auto { fix i assume asm: "n > 0" "i < n" "i > 0" and ih: "\k. k < i \ \j j \ A $$ (k, j) = 0" then have "\j. j i \ j \ A $$ (i, j) = 0" proof - have inter_part: "\b m e. (b::nat) < e \ b < m \ m < e \ {b.. {m..b m e f. (b::nat) < e \ b < m \ m < e \ (\k=b..k\{b..{m..b m e f. (b::nat) < e \ b < m \ m < e \ (\k=b..k=b..k=m..j. j < i \ A$$(i, j) = 0" by auto from tri upper_triangularD[of A j i] asm dim have zsi1: "\k. i < k \ k < n \ A$$(k, i) = 0" by auto have "(A * adjoint A)$$(i, i) = (\k=0.. = (\k=0..k=i.. = (\k=i.. = conjugate (A$$(i, i)) * A$$(i, i) + (\k=(Suc i)..k=(Suc i)..k=0.. = (\k=0..k=i.. = (\k=i.. = conjugate (A$$(i, i)) * A$$(i, i)" using asm zsi1 by (auto simp add: sum.atLeast_Suc_lessThan) finally have "(adjoint A * A)$$(i, i) = conjugate (A$$(i, i)) * A$$(i, i)" . with adaii eq0 have fsitoneq0: "(\k=(Suc i)..k. k i < k \ conjugate (A$$(i, k)) * A$$(i, k) = 0" by (rule sum_nonneg_0[of "{(Suc i)..k. k i A $$ (i, k) = 0" by auto with zsi0 show "\j. j i \ j \ A $$ (i, j) = 0" by (metis linorder_neqE_nat) qed } with case0 show "\i ia. 0 < n \ i < n \ ia < n \ (\k. k < ia \ \j j \ A $$ (k, j) = 0) \ \j j \ A $$ (ia, j) = 0" by auto qed qed lemma normal_complex_mat_has_spectral_decomposition: assumes A: "(A::complex mat) \ carrier_mat n n" and normal: "A * adjoint A = adjoint A * A" and c: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" and B: "unitary_schur_decomposition A es = (B,P,Q)" shows "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P" proof - have smw: "similar_mat_wit A B P (adjoint P)" and ut: "upper_triangular B" and uP: "unitary P" and dB: "diag_mat B = es" and "(Q = adjoint P)" using assms by (auto simp add: unitary_schur_decomposition) from smw have dimP: "P \ carrier_mat n n" and dimB: "B \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" unfolding similar_mat_wit_def using A by auto have dimaB: "adjoint B \ carrier_mat n n" using dimB by auto note dims = dimP dimB dimaP dimaB have "inverts_mat P (adjoint P)" using unitary_def uP dims by auto then have iaPP: "inverts_mat (adjoint P) P" using inverts_mat_symm using dims by auto have aPP: "adjoint P * P = 1\<^sub>m n" using dims iaPP unfolding inverts_mat_def by auto from smw have A: "A = P * B * (adjoint P)" unfolding similar_mat_wit_def Let_def by auto then have aA: "adjoint A = P * adjoint B * adjoint P" by (insert A dimP dimB dimaP, auto simp add: adjoint_mult[of _ n n _ n] adjoint_adjoint) have "A * adjoint A = (P * B * adjoint P) * (P * adjoint B * adjoint P)" using A aA by auto also have "\ = P * B * (adjoint P * P) * (adjoint B * adjoint P)" using dims by (mat_assoc n) also have "\ = P * B * 1\<^sub>m n * (adjoint B * adjoint P)" using dims aPP by (auto) also have "\ = P * B * adjoint B * adjoint P" using dims by (mat_assoc n) finally have "A * adjoint A = P * B * adjoint B * adjoint P". then have "adjoint P * (A * adjoint A) * P = (adjoint P * P) * B * adjoint B * (adjoint P * P)" using dims by (simp add: assoc_mult_mat[of _ n n _ n _ n]) also have "\ = 1\<^sub>m n * B * adjoint B * 1\<^sub>m n" using aPP by auto also have "\ = B * adjoint B" using dims by auto finally have eq0: "adjoint P * (A * adjoint A) * P = B * adjoint B". have "adjoint A * A = (P * adjoint B * adjoint P) * (P * B * adjoint P)" using A aA by auto also have "\ = P * adjoint B * (adjoint P * P) * (B * adjoint P)" using dims by (mat_assoc n) also have "\ = P * adjoint B * 1\<^sub>m n * (B * adjoint P)" using dims aPP by (auto) also have "\ = P * adjoint B * B * adjoint P" using dims by (mat_assoc n) finally have "adjoint A * A = P * adjoint B * B * adjoint P" by auto then have "adjoint P * (adjoint A * A) * P = (adjoint P * P) * adjoint B * B * (adjoint P * P)" using dims by (simp add: assoc_mult_mat[of _ n n _ n _ n]) also have "\ = 1\<^sub>m n * adjoint B * B * 1\<^sub>m n" using aPP by auto also have "\ = adjoint B * B" using dims by auto finally have eq1: "adjoint P * (adjoint A * A) * P = adjoint B * B". from normal have "adjoint P * (adjoint A * A) * P = adjoint P * (A * adjoint A) * P" by auto with eq0 eq1 have "B * adjoint B = adjoint B * B" by auto with ut dims have "diagonal_mat B" using normal_upper_triangular_matrix_is_diagonal by auto with smw uP dB show "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P" by auto qed lemma complex_mat_has_jordan_nf: fixes A :: "complex mat" assumes "A \ carrier_mat n n" shows "\n_as. jordan_nf A n_as" proof - have "\as. char_poly A = (\ a \ as. [:- a, 1:]) \ length as = n" using assms by (simp add: complex_mat_char_poly_factorizable) then show ?thesis using assms by (auto simp add: jordan_nf_iff_linear_factorization) qed lemma hermitian_is_normal: assumes "hermitian A" shows "A * adjoint A = adjoint A * A" using assms by (auto simp add: hermitian_def) lemma hermitian_eigenvalue_real: assumes dim: "(A::complex mat) \ carrier_mat n n" and hA: "hermitian A" and c: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" and B: "unitary_schur_decomposition A es = (B,P,Q)" shows "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P \ (\i < n. B$$(i, i) \ Reals)" proof - have normal: "A * adjoint A = adjoint A * A" using hA hermitian_is_normal by auto then have schur: "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P" using normal_complex_mat_has_spectral_decomposition[OF dim normal c B] by (simp) then have "similar_mat_wit A B P (adjoint P)" and uP: "unitary P" and dB: "diag_mat B = es" using assms by auto then have A: "A = P * B * (adjoint P)" and dimB: "B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" unfolding similar_mat_wit_def Let_def using dim by auto then have dimaB: "adjoint B \ carrier_mat n n" by auto have "adjoint A = adjoint (adjoint P) * adjoint (P * B)" apply (subst A) apply (subst adjoint_mult[of "P * B" n n "adjoint P" n]) apply (insert dimB dimP, auto) done also have "\ = P * adjoint (P * B)" by (auto simp add: adjoint_adjoint) also have "\ = P * (adjoint B * adjoint P)" using dimB dimP by (auto simp add: adjoint_mult) also have "\ = P * adjoint B * adjoint P" using dimB dimP by (subst assoc_mult_mat[symmetric, of P n n "adjoint B" n "adjoint P" n], auto) finally have aA: "adjoint A = P * adjoint B * adjoint P" . have "A = adjoint A" using hA hermitian_def[of A] by auto then have "P * B * adjoint P = P * adjoint B * adjoint P" using A aA by auto then have BaB: "B = adjoint B" using unitary_elim[OF dimB dimaB dimP] uP by auto { fix i assume "i < n" then have "B$$(i, i) = conjugate (B$$(i, i))" apply (subst BaB) by (insert dimB, simp add: adjoint_eval) then have "B$$(i, i) \ Reals" unfolding conjugate_complex_def using Reals_cnj_iff by auto } then have "\i Reals" by auto with schur show ?thesis by auto qed lemma hermitian_inner_prod_real: assumes dimA: "(A::complex mat) \ carrier_mat n n" and dimv: "v \ carrier_vec n" and hA: "hermitian A" shows "inner_prod v (A *\<^sub>v v) \ Reals" proof - obtain es where es: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" using complex_mat_char_poly_factorizable dimA by auto obtain B P Q where "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) then have "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P \ (\i < n. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real dimA es hA by auto then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and Bii: "\i. i < n \ B$$(i, i) \ Reals" and dimB: "B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" unfolding similar_mat_wit_def Let_def using dimA by auto define w where "w = (adjoint P) *\<^sub>v v" then have dimw: "w \ carrier_vec n" using dimaP dimv by auto from A have "inner_prod v (A *\<^sub>v v) = inner_prod v ((P * B * (adjoint P)) *\<^sub>v v)" by auto also have "\ = inner_prod v ((P * B) *\<^sub>v ((adjoint P) *\<^sub>v v))" using dimP dimB dimv by (subst assoc_mult_mat_vec[of _ n n "adjoint P" n], auto) also have "\ = inner_prod v (P *\<^sub>v (B *\<^sub>v ((adjoint P) *\<^sub>v v)))" using dimP dimB dimv dimaP by (subst assoc_mult_mat_vec[of _ n n "B" n], auto) also have "\ = inner_prod w (B *\<^sub>v w)" unfolding w_def apply (rule adjoint_def_alter[OF _ _ dimP]) apply (insert mult_mat_vec_carrier[OF dimB mult_mat_vec_carrier[OF dimaP dimv]], auto simp add: dimv) done also have "\ = (\i=0..j=0.. = (\i=0..v v) = (\i=0..i. i < n \ B$$(i, i) * conjugate (w$i) * w$i \ Reals" using Bii by (simp add: Reals_cnj_iff) then have "(\i=0.. Reals" by auto then show ?thesis using sum by auto qed lemma unit_vec_bracket: fixes A :: "complex mat" assumes dimA: "A \ carrier_mat n n" and i: "i < n" shows "inner_prod (unit_vec n i) (A *\<^sub>v (unit_vec n i)) = A$$(i, i)" proof - define w where "(w::complex vec) = unit_vec n i" have "A *\<^sub>v w = col A i" using i dimA w_def by auto then have 1: "inner_prod w (A *\<^sub>v w) = inner_prod w (col A i)" using w_def by auto have "conjugate w = w" unfolding w_def unit_vec_def conjugate_vec_def using i by auto then have 2: "inner_prod w (col A i) = A$$(i, i)" using i dimA w_def by auto from 1 2 show "inner_prod w (A *\<^sub>v w) = A$$(i, i)" by auto qed lemma spectral_decomposition_extract_diag: fixes P B :: "complex mat" assumes dimP: "P \ carrier_mat n n" and dimB: "B \ carrier_mat n n" and uP: "unitary P" and dB: "diagonal_mat B" and i: "i < n" shows "inner_prod (col P i) (P * B * (adjoint P) *\<^sub>v (col P i)) = B$$(i, i)" proof - have dimaP: "adjoint P\ carrier_mat n n" using dimP by auto have uaP: "unitary (adjoint P)" using unitary_adjoint uP dimP by auto then have "inverts_mat (adjoint P) P" by (simp add: unitary_def adjoint_adjoint) then have iv: "(adjoint P) * P = 1\<^sub>m n" using dimaP inverts_mat_def by auto define v where "v = col P i" then have dimv: "v \ carrier_vec n" using dimP by auto define w where "(w::complex vec) = unit_vec n i" then have dimw: "w \ carrier_vec n" by auto have BaPv: "B *\<^sub>v (adjoint P *\<^sub>v v) \ carrier_vec n" using dimB dimaP dimv by auto have "(adjoint P) *\<^sub>v v = (col (adjoint P * P) i)" by (simp add: col_mult2[OF dimaP dimP i, symmetric] v_def) then have aPv: "(adjoint P) *\<^sub>v v = w" by (auto simp add: iv i w_def) have "inner_prod v (P * B * (adjoint P) *\<^sub>v v) = inner_prod v ((P * B) *\<^sub>v ((adjoint P) *\<^sub>v v))" using dimP dimB dimv by (subst assoc_mult_mat_vec[of _ n n "adjoint P" n], auto) also have "\ = inner_prod v (P *\<^sub>v (B *\<^sub>v ((adjoint P) *\<^sub>v v)))" using dimP dimB dimv dimaP by (subst assoc_mult_mat_vec[of _ n n "B" n], auto) also have "\ = inner_prod (adjoint P *\<^sub>v v) (B *\<^sub>v (adjoint P *\<^sub>v v))" by (simp add: adjoint_def_alter[OF dimv BaPv dimP]) also have "\ = inner_prod w (B *\<^sub>v w)" using aPv by auto also have "\ = B$$(i, i)" using w_def unit_vec_bracket dimB i by auto finally show "inner_prod v (P * B * (adjoint P) *\<^sub>v v) = B$$(i, i)". qed lemma hermitian_inner_prod_zero: fixes A :: "complex mat" assumes dimA: "A \ carrier_mat n n" and hA: "hermitian A" and zero: "\v\carrier_vec n. inner_prod v (A *\<^sub>v v) = 0" shows "A = 0\<^sub>m n n" proof - obtain es where es: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" using complex_mat_char_poly_factorizable dimA by auto obtain B P Q where "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) then have "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P \ (\i < n. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real dimA es hA by auto then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and Bii: "\i. i < n \ B$$(i, i) \ Reals" and dimB: "B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" and uP: "unitary P" unfolding similar_mat_wit_def Let_def unitary_def using dimA by auto then have uaP: "unitary (adjoint P)" using unitary_adjoint by auto then have "inverts_mat (adjoint P) P" by (simp add: unitary_def adjoint_adjoint) then have iv: "adjoint P * P = 1\<^sub>m n" using dimaP inverts_mat_def by auto have "B = 0\<^sub>m n n" proof- { fix i assume i: "i < n" define v where "v = col P i" then have dimv: "v \ carrier_vec n" using v_def dimP by auto have "inner_prod v (A *\<^sub>v v) = B$$(i, i)" unfolding A v_def using spectral_decomposition_extract_diag[OF dimP dimB uP dB i] by auto moreover have "inner_prod v (A *\<^sub>v v) = 0" using dimv zero by auto ultimately have "B$$(i, i) = 0" by auto } note zB = this show "B = 0\<^sub>m n n" by (insert zB dB dimB, rule eq_matI, auto simp add: diagonal_mat_def) qed then show "A = 0\<^sub>m n n" using A dimB dimP dimaP by auto qed lemma complex_mat_decomposition_to_hermitian: fixes A :: "complex mat" assumes dim: "A \ carrier_mat n n" shows "\B C. hermitian B \ hermitian C \ A = B + \ \\<^sub>m C \ B \ carrier_mat n n \ C \ carrier_mat n n" proof - obtain B C where B: "B = (1 / 2) \\<^sub>m (A + adjoint A)" and C: "C = (-\ / 2) \\<^sub>m (A - adjoint A)" by auto then have dimB: "B \ carrier_mat n n" and dimC: "C \ carrier_mat n n" using dim by auto have "hermitian B" unfolding B hermitian_def using dim by (auto simp add: adjoint_eval) moreover have "hermitian C" unfolding C hermitian_def using dim apply (subst eq_matI) apply (auto simp add: adjoint_eval algebra_simps) done moreover have "A = B + \ \\<^sub>m C" using dim B C apply (subst eq_matI) apply (auto simp add: adjoint_eval algebra_simps) done ultimately show ?thesis using dimB dimC by auto qed subsection \Outer product\ definition outer_prod :: "'a::conjugatable_field vec \ 'a vec \ 'a mat" where "outer_prod v w = mat (dim_vec v) 1 (\(i, j). v $ i) * mat 1 (dim_vec w) (\(i, j). (conjugate w) $ j)" lemma outer_prod_dim[simp]: fixes v w :: "'a::conjugatable_field vec" assumes v: "v \ carrier_vec n" and w: "w \ carrier_vec m" shows "outer_prod v w \ carrier_mat n m" unfolding outer_prod_def using assms mat_of_cols_carrier mat_of_rows_carrier by auto lemma mat_of_vec_mult_eq_scalar_prod: fixes v w :: "'a::conjugatable_field vec" assumes "v \ carrier_vec n" and "w \ carrier_vec n" shows "mat 1 (dim_vec v) (\(i, j). (conjugate v) $ j) * mat (dim_vec w) 1 (\(i, j). w $ i) = mat 1 1 (\k. inner_prod v w)" apply (rule eq_matI) using assms apply (simp add: scalar_prod_def) apply (rule sum.cong) by auto lemma one_dim_mat_mult_is_scale: fixes A B :: "('a::conjugatable_field mat)" assumes "B \ carrier_mat 1 n" shows "(mat 1 1 (\k. a)) * B = a \\<^sub>m B" apply (rule eq_matI) using assms by (auto simp add: scalar_prod_def) lemma outer_prod_mult_outer_prod: fixes a b c d :: "'a::conjugatable_field vec" assumes a: "a \ carrier_vec d1" and b: "b \ carrier_vec d2" and c: "c \ carrier_vec d2" and d: "d \ carrier_vec d3" shows "outer_prod a b * outer_prod c d = inner_prod b c \\<^sub>m outer_prod a d" proof - let ?ma = "mat (dim_vec a) 1 (\(i, j). a $ i)" let ?mb = "mat 1 (dim_vec b) (\(i, j). (conjugate b) $ j)" let ?mc = "mat (dim_vec c) 1 (\(i, j). c $ i)" let ?md = "mat 1 (dim_vec d) (\(i, j). (conjugate d) $ j)" have "(?ma * ?mb) * (?mc * ?md) = ?ma * (?mb * (?mc * ?md))" apply (subst assoc_mult_mat[of "?ma" d1 1 "?mb" d2 "?mc * ?md" d3] ) using assms by auto also have "\ = ?ma * ((?mb * ?mc) * ?md)" apply (subst assoc_mult_mat[symmetric, of "?mb" 1 d2 "?mc" 1 "?md" d3]) using assms by auto also have "\ = ?ma * ((mat 1 1 (\k. inner_prod b c)) * ?md)" apply (subst mat_of_vec_mult_eq_scalar_prod[of b d2 c]) using assms by auto also have "\ = ?ma * (inner_prod b c \\<^sub>m ?md)" apply (subst one_dim_mat_mult_is_scale) using assms by auto also have "\ = (inner_prod b c) \\<^sub>m (?ma * ?md)" using assms by auto finally show ?thesis unfolding outer_prod_def by auto qed lemma index_outer_prod: fixes v w :: "'a::conjugatable_field vec" assumes v: "v \ carrier_vec n" and w: "w \ carrier_vec m" and ij: "i < n" "j < m" shows "(outer_prod v w)$$(i, j) = v $ i * conjugate (w $ j)" unfolding outer_prod_def using assms by (simp add: scalar_prod_def) lemma mat_of_vec_mult_vec: fixes a b c :: "'a::conjugatable_field vec" assumes a: "a \ carrier_vec d" and b: "b \ carrier_vec d" shows "mat 1 d (\(i, j). (conjugate a) $ j) *\<^sub>v b = vec 1 (\k. inner_prod a b)" apply (rule eq_vecI) apply (simp add: scalar_prod_def carrier_vecD[OF a] carrier_vecD[OF b]) apply (rule sum.cong) by auto lemma mat_of_vec_mult_one_dim_vec: fixes a b :: "'a::conjugatable_field vec" assumes a: "a \ carrier_vec d" shows "mat d 1 (\(i, j). a $ i) *\<^sub>v vec 1 (\k. c) = c \\<^sub>v a" apply (rule eq_vecI) by (auto simp add: scalar_prod_def carrier_vecD[OF a]) lemma outer_prod_mult_vec: fixes a b c :: "'a::conjugatable_field vec" assumes a: "a \ carrier_vec d1" and b: "b \ carrier_vec d2" and c: "c \ carrier_vec d2" shows "outer_prod a b *\<^sub>v c = inner_prod b c \\<^sub>v a" proof - have "outer_prod a b *\<^sub>v c = mat d1 1 (\(i, j). a $ i) * mat 1 d2 (\(i, j). (conjugate b) $ j) *\<^sub>v c" unfolding outer_prod_def using assms by auto also have "\ = mat d1 1 (\(i, j). a $ i) *\<^sub>v (mat 1 d2 (\(i, j). (conjugate b) $ j) *\<^sub>v c)" apply (subst assoc_mult_mat_vec) using assms by auto also have "\ = mat d1 1 (\(i, j). a $ i) *\<^sub>v vec 1 (\k. inner_prod b c)" using mat_of_vec_mult_vec[of b] assms by auto also have "\ = inner_prod b c \\<^sub>v a" using mat_of_vec_mult_one_dim_vec assms by auto finally show ?thesis by auto qed lemma trace_outer_prod_right: fixes A :: "'a::conjugatable_field mat" and v w :: "'a vec" assumes A: "A \ carrier_mat n n" and v: "v \ carrier_vec n" and w: "w \ carrier_vec n" shows "trace (A * outer_prod v w) = inner_prod w (A *\<^sub>v v)" (is "?lhs = ?rhs") proof - define B where "B = outer_prod v w" then have B: "B \ carrier_mat n n" using assms by auto have "trace(A * B) = (\i = 0..j = 0.. = (\i = 0..j = 0..i = 0..j = 0..i = 0..j = 0.. carrier_vec n" and w: "w \ carrier_vec n" shows "trace (outer_prod v w) = inner_prod w v" (is "?lhs = ?rhs") proof - have "(1\<^sub>m n) * (outer_prod v w) = outer_prod v w" apply (subst left_mult_one_mat) using outer_prod_dim assms by auto moreover have "1\<^sub>m n *\<^sub>v v = v" using assms by auto ultimately show ?thesis using trace_outer_prod_right[of "1\<^sub>m n" n v w] assms by auto qed lemma inner_prod_outer_prod: fixes a b c d :: "'a::conjugatable_field vec" assumes a: "a \ carrier_vec n" and b: "b \ carrier_vec n" and c: "c \ carrier_vec m" and d: "d \ carrier_vec m" shows "inner_prod a (outer_prod b c *\<^sub>v d) = inner_prod a b * inner_prod c d" (is "?lhs = ?rhs") proof - define P where "P = outer_prod b c" then have dimP: "P \ carrier_mat n m" using assms by auto have "inner_prod a (P *\<^sub>v d) = (\i=0..j=0.. = (\i=0..j=0..i=0..j=0..i=0..j=0.. = (\i=0..j=0..Semi-definite matrices\ definition positive :: "complex mat \ bool" where "positive A \ A \ carrier_mat (dim_col A) (dim_col A) \ (\v. dim_vec v = dim_col A \ inner_prod v (A *\<^sub>v v) \ 0)" lemma positive_iff_normalized_vec: "positive A \ A \ carrier_mat (dim_col A) (dim_col A) \ (\v. (dim_vec v = dim_col A \ vec_norm v = 1) \ inner_prod v (A *\<^sub>v v) \ 0)" proof (rule) assume "positive A" then show "A \ carrier_mat (dim_col A) (dim_col A) \ (\v. dim_vec v = dim_col A \ vec_norm v = 1 \ 0 \ inner_prod v (A *\<^sub>v v))" unfolding positive_def by auto next define n where "n = dim_col A" assume "A \ carrier_mat (dim_col A) (dim_col A) \ (\v. dim_vec v = dim_col A \ vec_norm v = 1 \ 0 \ inner_prod v (A *\<^sub>v v))" then have A: "A \ carrier_mat (dim_col A) (dim_col A)" and geq0: "\v. dim_vec v = dim_col A \ vec_norm v = 1 \ 0 \ inner_prod v (A *\<^sub>v v)" by auto then have dimA: "A \ carrier_mat n n" using n_def[symmetric] by auto { fix v assume dimv: "(v::complex vec) \ carrier_vec n" have "0 \ inner_prod v (A *\<^sub>v v)" proof (cases "v = 0\<^sub>v n") case True then show "0 \ inner_prod v (A *\<^sub>v v)" using dimA by auto next case False then have 1: "vec_norm v > 0" using vec_norm_ge_0 dimv by auto then have cnv: "cnj (vec_norm v) = vec_norm v" using Reals_cnj_iff complex_is_Real_iff less_complex_def by auto define w where "w = vec_normalize v" then have dimw: "w \ carrier_vec n" using dimv by auto have nvw: "v = vec_norm v \\<^sub>v w" using w_def vec_eq_norm_smult_normalized by auto have "vec_norm w = 1" using normalized_vec_norm[OF dimv False] vec_norm_def w_def by auto then have 2: "0 \ inner_prod w (A *\<^sub>v w)" using geq0 dimw dimA by auto have "inner_prod v (A *\<^sub>v v) = vec_norm v * vec_norm v * inner_prod w (A *\<^sub>v w)" using dimA dimv dimw apply (subst (1 2) nvw) apply (subst mult_mat_vec, simp, simp) apply (subst scalar_prod_smult_left[of "(A *\<^sub>v w)" "conjugate (vec_norm v \\<^sub>v w)" "vec_norm v"], simp) apply (simp add: conjugate_smult_vec cnv) done also have "\ \ 0" using 1 2 by auto finally show "0 \ inner_prod v (A *\<^sub>v v)" by auto qed } then have geq: "\v. dim_vec v = dim_col A \ 0 \ inner_prod v (A *\<^sub>v v)" using dimA by auto show "positive A" unfolding positive_def by (rule, simp add: A, rule geq) qed lemma positive_is_hermitian: fixes A :: "complex mat" assumes pA: "positive A" shows "hermitian A" proof - define n where "n = dim_col A" then have dimA: "A \ carrier_mat n n" using positive_def pA by auto obtain B C where B: "hermitian B" and C: "hermitian C" and A: "A = B + \ \\<^sub>m C" and dimB: "B \ carrier_mat n n" and dimC: "C \ carrier_mat n n" and dimiC: "\ \\<^sub>m C \ carrier_mat n n" using complex_mat_decomposition_to_hermitian[OF dimA] by auto { fix v :: "complex vec" assume dimv: "v \ carrier_vec n" have dimvA: "dim_vec v = dim_col A" using dimv dimA by auto have "inner_prod v (A *\<^sub>v v) = inner_prod v (B *\<^sub>v v) + inner_prod v ((\ \\<^sub>m C) *\<^sub>v v)" unfolding A using dimB dimiC dimv by (simp add: add_mult_distrib_mat_vec inner_prod_distrib_right) moreover have "inner_prod v ((\ \\<^sub>m C) *\<^sub>v v) = \ * inner_prod v (C *\<^sub>v v)" using dimv dimC apply (simp add: scalar_prod_def sum_distrib_left cong: sum.cong) apply (rule sum.cong, auto) done ultimately have ABC: "inner_prod v (A *\<^sub>v v) = inner_prod v (B *\<^sub>v v) + \ * inner_prod v (C *\<^sub>v v)" by auto moreover have "inner_prod v (B *\<^sub>v v) \ Reals" using B dimB dimv hermitian_inner_prod_real by auto moreover have "inner_prod v (C *\<^sub>v v) \ Reals" using C dimC dimv hermitian_inner_prod_real by auto moreover have "inner_prod v (A *\<^sub>v v) \ Reals" using pA unfolding positive_def apply (rule) apply (fold n_def) apply (simp add: complex_is_Real_iff[of "inner_prod v (A *\<^sub>v v)"]) apply (auto simp add: dimvA less_complex_def less_eq_complex_def) done ultimately have "inner_prod v (C *\<^sub>v v) = 0" using of_real_Re by fastforce } then have "C = 0\<^sub>m n n" using hermitian_inner_prod_zero dimC C by auto then have "A = B" using A dimC dimB by auto then show "hermitian A" using B by auto qed lemma positive_eigenvalue_positive: assumes dimA: "(A::complex mat) \ carrier_mat n n" and pA: "positive A" and c: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" and B: "unitary_schur_decomposition A es = (B,P,Q)" shows "\i. i < n \ B$$(i, i) \ 0" proof - have hA: "hermitian A" using positive_is_hermitian pA by auto have "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P \ (\i < n. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real dimA hA B c by auto then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and Bii: "\i. i < n \ B$$(i, i) \ Reals" and dimB: "B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" and uP: "unitary P" unfolding similar_mat_wit_def Let_def unitary_def using dimA by auto { fix i assume i: "i < n" define v where "v = col P i" then have dimv: "v \ carrier_vec n" using v_def dimP by auto have "inner_prod v (A *\<^sub>v v) = B$$(i, i)" unfolding A v_def using spectral_decomposition_extract_diag[OF dimP dimB uP dB i] by auto moreover have "inner_prod v (A *\<^sub>v v) \ 0" using dimv pA dimA positive_def by auto ultimately show "B$$(i, i) \ 0" by auto } qed lemma diag_mat_mult_diag_mat: fixes B D :: "'a::semiring_0 mat" assumes dimB: "B \ carrier_mat n n" and dimD: "D \ carrier_mat n n" and dB: "diagonal_mat B" and dD: "diagonal_mat D" shows "B * D = mat n n (\(i,j). (if i = j then (B$$(i, i)) * (D$$(i, i)) else 0))" proof(rule eq_matI, auto) have Bij: "\x y. x < n \ y < n \ x \ y \ B$$(x, y) = 0" using dB diagonal_mat_def dimB by auto have Dij: "\x y. x < n \ y < n \ x \ y \ D$$(x, y) = 0" using dD diagonal_mat_def dimD by auto { fix i j assume ij: "i < n" "j < n" have "(B * D) $$ (i, j) = (\k=0.. = B$$(i, i) * D$$(i, j)" apply (simp add: sum.remove[of _i] ij) apply (simp add: Bij Dij ij) done finally have "(B * D) $$ (i, j) = B$$(i, i) * D$$(i, j)". } note BDij = this from BDij show "\j. j < n \ (B * D) $$ (j, j) = B $$ (j, j) * D $$ (j, j)" by auto from BDij show "\i j. i < n \ j < n \ i \ j \ (B * D) $$ (i, j) = 0" using Bij Dij by auto from assms show "dim_row B = n" "dim_col D = n" by auto qed lemma positive_only_if_decomp: assumes dimA: "A \ carrier_mat n n" and pA: "positive A" shows "\M \ carrier_mat n n. M * adjoint M = A" proof - from pA have hA: "hermitian A" using positive_is_hermitian by auto obtain es where es: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" using complex_mat_char_poly_factorizable dimA by auto obtain B P Q where schur: "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) then have "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P \ (\i < n. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real dimA es hA by auto then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and Bii: "\i. i < n \ B$$(i, i) \ Reals" and dimB: "B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" unfolding similar_mat_wit_def Let_def using dimA by auto have Bii: "\i. i < n \ B$$(i, i) \ 0" using pA dimA es schur positive_eigenvalue_positive by auto define D where "D = mat n n (\(i, j). (if (i = j) then csqrt (B$$(i, i)) else 0))" then have dimD: "D \ carrier_mat n n" and dimaD: "adjoint D \ carrier_mat n n" using dimB by auto have dD: "diagonal_mat D" using dB D_def unfolding diagonal_mat_def by auto then have daD: "diagonal_mat (adjoint D)" by (simp add: adjoint_eval diagonal_mat_def) have Dii: "\i. i < n \ D$$(i, i) = csqrt (B$$(i, i))" using dimD D_def by auto { fix i assume i: "i < n" define c where "c = csqrt (B$$(i, i))" have c: "c \ 0" using Bii i c_def by (auto simp: less_complex_def less_eq_complex_def) then have "conjugate c = c" using Reals_cnj_iff complex_is_Real_iff unfolding less_complex_def less_eq_complex_def by auto then have "c * cnj c = B$$(i, i)" using c_def c unfolding conjugate_complex_def by (metis power2_csqrt power2_eq_square) } note cBii = this have "D * adjoint D = mat n n (\(i,j). (if (i = j) then B$$(i, i) else 0))" apply (simp add: diag_mat_mult_diag_mat[OF dimD dimaD dD daD]) apply (rule eq_matI, auto simp add: D_def adjoint_eval cBii) done also have "\ = B" using dimB dB[unfolded diagonal_mat_def] by auto finally have DaDB: "D * adjoint D = B". define M where "M = P * D" then have dimM: "M \ carrier_mat n n" using dimP dimD by auto have "M * adjoint M = (P * D) * (adjoint D * adjoint P)" using M_def adjoint_mult[OF dimP dimD] by auto also have "\ = P * (D * adjoint D) * (adjoint P)" using dimP dimD by (mat_assoc n) also have "\ = P * B * (adjoint P)" using DaDB by auto finally have "M * adjoint M = A" using A by auto with dimM show "\M \ carrier_mat n n. M * adjoint M = A" by auto qed lemma positive_if_decomp: assumes dimA: "A \ carrier_mat n n" and "\M. M * adjoint M = A" shows "positive A" proof - from assms obtain M where M: "M * adjoint M = A" by auto define m where "m = dim_col M" have dimM: "M \ carrier_mat n m" using M dimA m_def by auto { fix v assume dimv: "(v::complex vec) \ carrier_vec n" have dimaM: "adjoint M \ carrier_mat m n" using dimM by auto have dimaMv: "(adjoint M) *\<^sub>v v \ carrier_vec m" using dimaM dimv by auto have "inner_prod v (A *\<^sub>v v) = inner_prod v (M * adjoint M *\<^sub>v v)" using M by auto also have "\ = inner_prod v (M *\<^sub>v (adjoint M *\<^sub>v v))" using assoc_mult_mat_vec dimM dimaM dimv by auto also have "\ = inner_prod (adjoint M *\<^sub>v v) (adjoint M *\<^sub>v v)" using adjoint_def_alter[OF dimv dimaMv dimM] by auto also have "\ \ 0" using self_cscalar_prod_geq_0 by auto finally have "inner_prod v (A *\<^sub>v v) \ 0". } note geq0 = this from dimA geq0 show "positive A" using positive_def by auto qed lemma positive_iff_decomp: assumes dimA: "A \ carrier_mat n n" shows "positive A \ (\M\carrier_mat n n. M * adjoint M = A)" proof assume pA: "positive A" then show "\M\carrier_mat n n. M * adjoint M = A" using positive_only_if_decomp assms by auto next assume "\M\carrier_mat n n. M * adjoint M = A" then obtain M where M: "M * adjoint M = A" by auto then show "positive A" using M positive_if_decomp assms by auto qed lemma positive_dim_eq: assumes "positive A" shows "dim_row A = dim_col A" using carrier_matD(1)[of A "dim_col A" "dim_col A"] assms[unfolded positive_def] by simp lemma positive_zero: "positive (0\<^sub>m n n)" by (simp add: positive_def zero_mat_def mult_mat_vec_def scalar_prod_def) lemma positive_one: "positive (1\<^sub>m n)" proof (rule positive_if_decomp) show "1\<^sub>m n \ carrier_mat n n" by auto have "adjoint (1\<^sub>m n) = 1\<^sub>m n" using hermitian_one hermitian_def by auto then have "1\<^sub>m n * adjoint (1\<^sub>m n) = 1\<^sub>m n" by auto then show "\M. M * adjoint M = 1\<^sub>m n" by fastforce qed lemma positive_antisym: assumes pA: "positive A" and pnA: "positive (-A)" shows "A = 0\<^sub>m (dim_col A) (dim_col A)" proof - define n where "n = dim_col A" from pA have dimA: "A \ carrier_mat n n" and dimnA: "-A \ carrier_mat n n" using positive_def n_def by auto from pA have hA: "hermitian A" using positive_is_hermitian by auto obtain es where es: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" using complex_mat_char_poly_factorizable dimA by auto obtain B P Q where schur: "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) then have "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ unitary P" using hermitian_eigenvalue_real dimA es hA by auto then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and uP: "unitary P" and dimB: "B \ carrier_mat n n" and dimnB: "-B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" unfolding similar_mat_wit_def Let_def using dimA by auto from es schur have geq0: "\i. i < n \ B$$(i, i) \ 0" using positive_eigenvalue_positive dimA pA by auto from A have nA: "-A = P * (-B) * (adjoint P)" using mult_smult_assoc_mat dimB dimP dimaP by auto from dB have dnB: "diagonal_mat (-B)" by (simp add: diagonal_mat_def) { fix i assume i: "i < n" define v where "v = col P i" then have dimv: "v \ carrier_vec n" using v_def dimP by auto have "inner_prod v ((-A) *\<^sub>v v) = (-B)$$(i, i)" unfolding nA v_def using spectral_decomposition_extract_diag[OF dimP dimnB uP dnB i] by auto moreover have "inner_prod v ((-A) *\<^sub>v v) \ 0" using dimv pnA dimnA positive_def by auto ultimately have "B$$(i, i) \ 0" using dimB i by auto moreover have "B$$(i, i) \ 0" using i geq0 by auto - ultimately have "B$$(i, i) = 0" by (metis no_atp(10)) + ultimately have "B$$(i, i) = 0" + by (metis antisym) } then have "B = 0\<^sub>m n n" using dimB dB[unfolded diagonal_mat_def] by (subst eq_matI, auto) then show "A = 0\<^sub>m n n" using A dimB dimP dimaP by auto qed lemma positive_add: assumes pA: "positive A" and pB: "positive B" and dimA: "A \ carrier_mat n n" and dimB: "B \ carrier_mat n n" shows "positive (A + B)" unfolding positive_def proof have dimApB: "A + B \ carrier_mat n n" using dimA dimB by auto then show "A + B \ carrier_mat (dim_col (A + B)) (dim_col (A + B))" using carrier_matD[of "A+B"] by auto { fix v assume dimv: "(v::complex vec) \ carrier_vec n" have 1: "inner_prod v (A *\<^sub>v v) \ 0" using dimv pA[unfolded positive_def] dimA by auto have 2: "inner_prod v (B *\<^sub>v v) \ 0" using dimv pB[unfolded positive_def] dimB by auto have "inner_prod v ((A + B) *\<^sub>v v) = inner_prod v (A *\<^sub>v v) + inner_prod v (B *\<^sub>v v)" using dimA dimB dimv by (simp add: add_mult_distrib_mat_vec inner_prod_distrib_right) also have "\ \ 0" using 1 2 by auto finally have "inner_prod v ((A + B) *\<^sub>v v) \ 0". } note geq0 = this then have "\v. dim_vec v = n \ 0 \ inner_prod v ((A + B) *\<^sub>v v)" by auto then show "\v. dim_vec v = dim_col (A + B) \ 0 \ inner_prod v ((A + B) *\<^sub>v v)" using dimApB by auto qed lemma positive_trace: assumes "A \ carrier_mat n n" and "positive A" shows "trace A \ 0" using assms positive_iff_decomp trace_adjoint_positive by auto lemma positive_close_under_left_right_mult_adjoint: fixes M A :: "complex mat" assumes dM: "M \ carrier_mat n n" and dA: "A \ carrier_mat n n" and pA: "positive A" shows "positive (M * A * adjoint M)" unfolding positive_def proof (rule, simp add: mult_carrier_mat[OF mult_carrier_mat[OF dM dA] adjoint_dim[OF dM]] carrier_matD[OF dM], rule, rule) have daM: "adjoint M \ carrier_mat n n" using dM by auto fix v::"complex vec" assume "dim_vec v = dim_col (M * A * adjoint M)" then have dv: "v \ carrier_vec n" using assms by auto then have "adjoint M *\<^sub>v v \ carrier_vec n" using daM by auto have assoc: "M * A * adjoint M *\<^sub>v v = M *\<^sub>v (A *\<^sub>v (adjoint M *\<^sub>v v))" using dA dM daM dv by (auto simp add: assoc_mult_mat_vec[of _ n n _ n]) have "inner_prod v (M * A * adjoint M *\<^sub>v v) = inner_prod (adjoint M *\<^sub>v v) (A *\<^sub>v (adjoint M *\<^sub>v v))" apply (subst assoc) apply (subst adjoint_def_alter[where ?A = "M"]) by (auto simp add: dv dA daM dM carrier_matD[OF dM] mult_mat_vec_carrier[of _ n n]) also have "\ \ 0" using dA dv daM pA positive_def by auto finally show "inner_prod v (M * A * adjoint M *\<^sub>v v) \ 0" by auto qed lemma positive_same_outer_prod: fixes v w :: "complex vec" assumes v: "v \ carrier_vec n" shows "positive (outer_prod v v)" proof - have d1: "adjoint (mat (dim_vec v) 1 (\(i, j). v $ i)) \ carrier_mat 1 n" using assms by auto have d2: "mat 1 (dim_vec v) (\(i, y). conjugate v $ y) \ carrier_mat 1 n" using assms by auto have dv: "dim_vec v = n" using assms by auto have "mat 1 (dim_vec v) (\(i, y). conjugate v $ y) = adjoint (mat (dim_vec v) 1 (\(i, j). v $ i))" (is "?r = adjoint ?l") apply (rule eq_matI) subgoal for i j by (simp add: dv adjoint_eval) using d1 d2 by auto then have "outer_prod v v = ?l * adjoint ?l" unfolding outer_prod_def by auto then have "\M. M * adjoint M = outer_prod v v" by auto then show "positive (outer_prod v v)" using positive_if_decomp[OF outer_prod_dim[OF v v]] by auto qed lemma smult_smult_mat: fixes k :: complex and l :: complex assumes "A \ carrier_mat nr n" shows "k \\<^sub>m (l \\<^sub>m A) = (k * l) \\<^sub>m A" by auto lemma positive_smult: assumes "A \ carrier_mat n n" and "positive A" and "c \ 0" shows "positive (c \\<^sub>m A)" proof - have sc: "csqrt c \ 0" using assms(3) by (fastforce simp: less_eq_complex_def) obtain M where dimM: "M \ carrier_mat n n" and A: "M * adjoint M = A" using assms(1-2) positive_iff_decomp by auto have "c \\<^sub>m A = c \\<^sub>m (M * adjoint M)" using A by auto have ccsq: "conjugate (csqrt c) = (csqrt c)" using sc Reals_cnj_iff[of "csqrt c"] complex_is_Real_iff by (auto simp: less_eq_complex_def) have MM: "(M * adjoint M) \ carrier_mat n n" using A assms by fastforce have leftd: "c \\<^sub>m (M * adjoint M) \ carrier_mat n n" using A assms by fastforce have rightd: "(csqrt c \\<^sub>m M) * (adjoint (csqrt c \\<^sub>m M))\ carrier_mat n n" using A assms by fastforce have "(csqrt c \\<^sub>m M) * (adjoint (csqrt c \\<^sub>m M)) = (csqrt c \\<^sub>m M) * ((conjugate (csqrt c)) \\<^sub>m adjoint M)" using adjoint_scale assms(1) by (metis adjoint_scale) also have "\ = (csqrt c \\<^sub>m M) * (csqrt c \\<^sub>m adjoint M)" using sc ccsq by fastforce also have "\ = csqrt c \\<^sub>m (M * (csqrt c \\<^sub>m adjoint M))" using mult_smult_assoc_mat index_smult_mat(2,3) by fastforce also have "\ = csqrt c \\<^sub>m ((csqrt c) \\<^sub>m (M * adjoint M))" using mult_smult_distrib by fastforce also have "\ = c \\<^sub>m (M * adjoint M)" using smult_smult_mat[of "M * adjoint M" n n "(csqrt c)" "(csqrt c)"] MM sc by (metis power2_csqrt power2_eq_square ) also have "\ = c \\<^sub>m A" using A by auto finally have "(csqrt c \\<^sub>m M) * (adjoint (csqrt c \\<^sub>m M)) = c \\<^sub>m A" by auto moreover have "c \\<^sub>m A \ carrier_mat n n" using assms(1) by auto moreover have "csqrt c \\<^sub>m M \ carrier_mat n n" using dimM by auto ultimately show ?thesis using positive_iff_decomp by auto qed text \Version of previous theorem for real numbers\ lemma positive_scale: fixes c :: real assumes "A \ carrier_mat n n" and "positive A" and "c \ 0" shows "positive (c \\<^sub>m A)" apply (rule positive_smult) using assms by (auto simp: less_eq_complex_def) subsection \L\"{o}wner partial order\ definition lowner_le :: "complex mat \ complex mat \ bool" (infix "\\<^sub>L" 50) where "A \\<^sub>L B \ dim_row A = dim_row B \ dim_col A = dim_col B \ positive (B - A)" lemma lowner_le_refl: assumes "A \ carrier_mat n n" shows "A \\<^sub>L A" unfolding lowner_le_def apply (simp add: minus_r_inv_mat[OF assms]) by (rule positive_zero) lemma lowner_le_antisym: assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and L1: "A \\<^sub>L B" and L2: "B \\<^sub>L A" shows "A = B" proof - from L1 have P1: "positive (B - A)" by (simp add: lowner_le_def) from L2 have P2: "positive (A - B)" by (simp add: lowner_le_def) have "A - B = - (B - A)" using A B by auto then have P3: "positive (- (B - A))" using P2 by auto have BA: "B - A \ carrier_mat n n" using A B by auto have "B - A = 0\<^sub>m n n" using BA by (subst positive_antisym[OF P1 P3], auto) then have "B + (-A) + A = 0\<^sub>m n n + A" using A B minus_add_uminus_mat[OF B A] by auto then have "B + (-A + A) = 0\<^sub>m n n + A" using A B by auto then show "A = B" using A B BA uminus_l_inv_mat[OF A] by auto qed lemma lowner_le_inner_prod_le: fixes A B :: "complex mat" and v :: "complex vec" assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and v: "v \ carrier_vec n" and "A \\<^sub>L B" shows "inner_prod v (A *\<^sub>v v) \ inner_prod v (B *\<^sub>v v)" proof - from assms have "positive (B-A)" by (auto simp add: lowner_le_def) with assms have geq: "inner_prod v ((B-A) *\<^sub>v v) \ 0" unfolding positive_def by auto have "inner_prod v ((B-A) *\<^sub>v v) = inner_prod v (B *\<^sub>v v) - inner_prod v (A *\<^sub>v v)" unfolding minus_add_uminus_mat[OF B A] by (subst add_mult_distrib_mat_vec[OF B _ v], insert A B v, auto simp add: inner_prod_distrib_right[OF v]) then show ?thesis using geq by auto qed lemma lowner_le_trans: fixes A B C :: "complex mat" assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and C: "C \ carrier_mat n n" and L1: "A \\<^sub>L B" and L2: "B \\<^sub>L C" shows "A \\<^sub>L C" unfolding lowner_le_def proof (auto simp add: carrier_matD[OF A] carrier_matD[OF C]) have dim: "C - A \ carrier_mat n n" using A C by auto { fix v assume v: "(v::complex vec) \ carrier_vec n" from L1 have "inner_prod v (A *\<^sub>v v) \ inner_prod v (B *\<^sub>v v)" using lowner_le_inner_prod_le A B v by auto also from L2 have "\ \ inner_prod v (C *\<^sub>v v)" using lowner_le_inner_prod_le B C v by auto finally have "inner_prod v (A *\<^sub>v v) \ inner_prod v (C *\<^sub>v v)". then have "inner_prod v (C *\<^sub>v v) - inner_prod v (A *\<^sub>v v) \ 0" by auto then have "inner_prod v ((C - A) *\<^sub>v v) \ 0" using A C v apply (subst minus_add_uminus_mat[OF C A]) apply (subst add_mult_distrib_mat_vec[OF C _ v], simp) apply (simp add: inner_prod_distrib_right[OF v]) done } note leq = this show "positive (C - A)" unfolding positive_def apply (rule, simp add: carrier_matD[OF A] dim) apply (subst carrier_matD[OF dim], insert leq, auto) done qed lemma lowner_le_imp_trace_le: assumes "A \ carrier_mat n n" and "B \ carrier_mat n n" and "A \\<^sub>L B" shows "trace A \ trace B" proof - have "positive (B - A)" using assms lowner_le_def by auto moreover have "B - A \ carrier_mat n n" using assms by auto ultimately have "trace (B - A) \ 0" using positive_trace by auto moreover have "trace (B - A) = trace B - trace A" using trace_minus_linear assms by auto ultimately have "trace B - trace A \ 0" by auto then show "trace A \ trace B" by auto qed lemma lowner_le_add: assumes "A \ carrier_mat n n" "B \ carrier_mat n n" "C \ carrier_mat n n" "D \ carrier_mat n n" and "A \\<^sub>L B" "C \\<^sub>L D" shows "A + C \\<^sub>L B + D" proof - have "B + D - (A + C) = B - A + (D - C) " using assms by auto then have "positive (B + D - (A + C))" using assms unfolding lowner_le_def using positive_add by (metis minus_carrier_mat) then show "A + C \\<^sub>L B + D" unfolding lowner_le_def using assms by fastforce qed lemma lowner_le_swap: assumes "A \ carrier_mat n n" "B \ carrier_mat n n" and "A \\<^sub>L B" shows "-B \\<^sub>L -A" proof - have "positive (B - A)" using assms lowner_le_def by fastforce moreover have "B - A = (-A) - (-B)" using assms by fastforce ultimately have "positive ((-A) - (-B))" by auto then show ?thesis using lowner_le_def assms by fastforce qed lemma lowner_le_minus: assumes "A \ carrier_mat n n" "B \ carrier_mat n n" "C \ carrier_mat n n" "D \ carrier_mat n n" and "A \\<^sub>L B" "C \\<^sub>L D" shows "A - D \\<^sub>L B - C" proof - have "positive (D - C)" using assms lowner_le_def by auto then have "-D \\<^sub>L -C" using lowner_le_swap assms by auto then have "A + (-D) \\<^sub>L B + (-C)" using lowner_le_add[of "A" n "B"] assms by auto moreover have "A + (-D) = A - D" and "B + (-C) = B - C" by auto ultimately show ?thesis by auto qed lemma outer_prod_le_one: assumes "v \ carrier_vec n" and "inner_prod v v \ 1" shows "outer_prod v v \\<^sub>L 1\<^sub>m n" proof - let ?o = "outer_prod v v" have do: "?o \ carrier_mat n n" using assms by auto { fix u :: "complex vec" assume "dim_vec u = n" then have du: "u \ carrier_vec n" by auto have r: "inner_prod u u \ Reals" apply (simp add: scalar_prod_def carrier_vecD[OF du]) using complex_In_mult_cnj_zero complex_is_Real_iff by blast have geq0: "inner_prod u u \ 0" using self_cscalar_prod_geq_0 by auto have "inner_prod u (?o *\<^sub>v u) = inner_prod u v * inner_prod v u" apply (subst inner_prod_outer_prod) using du assms by auto also have "\ \ inner_prod u u * inner_prod v v" using Cauchy_Schwarz_complex_vec du assms by auto also have "\ \ inner_prod u u" using assms(2) r geq0 by (simp add: mult_right_le_one_le less_eq_complex_def) finally have le: "inner_prod u (?o *\<^sub>v u) \ inner_prod u u". have "inner_prod u ((1\<^sub>m n - ?o) *\<^sub>v u) = inner_prod u ((1\<^sub>m n *\<^sub>v u) - ?o *\<^sub>v u)" apply (subst minus_mult_distrib_mat_vec) using do du by auto also have "\ = inner_prod u u - inner_prod u (?o *\<^sub>v u)" apply (subst inner_prod_minus_distrib_right) using du do by auto also have "\ \ 0" using le by auto finally have "inner_prod u ((1\<^sub>m n - ?o) *\<^sub>v u) \ 0" by auto } then have "positive (1\<^sub>m n - outer_prod v v)" unfolding positive_def using do by auto then show ?thesis unfolding lowner_le_def using do by auto qed lemma zero_lowner_le_positiveD: fixes A :: "complex mat" assumes dA: "A \ carrier_mat n n" and le: "0\<^sub>m n n \\<^sub>L A" shows "positive A" using assms unfolding lowner_le_def by (subgoal_tac "A - 0\<^sub>m n n = A", auto) lemma zero_lowner_le_positiveI: fixes A :: "complex mat" assumes dA: "A \ carrier_mat n n" and le: "positive A" shows "0\<^sub>m n n \\<^sub>L A" using assms unfolding lowner_le_def by (subgoal_tac "A - 0\<^sub>m n n = A", auto) lemma lowner_le_trans_positiveI: fixes A B :: "complex mat" assumes dA: "A \ carrier_mat n n" and pA: "positive A" and le: "A \\<^sub>L B" shows "positive B" proof - have dB: "B \ carrier_mat n n" using le dA lowner_le_def by auto have "0\<^sub>m n n \\<^sub>L A" using zero_lowner_le_positiveI dA pA by auto then have "0\<^sub>m n n \\<^sub>L B" using dA dB le by (simp add: lowner_le_trans[of _ n A B]) then show ?thesis using dB zero_lowner_le_positiveD by auto qed lemma lowner_le_keep_under_measurement: fixes M A B :: "complex mat" assumes dM: "M \ carrier_mat n n" and dA: "A \ carrier_mat n n" and dB: "B \ carrier_mat n n" and le: "A \\<^sub>L B" shows "adjoint M * A * M \\<^sub>L adjoint M * B * M" unfolding lowner_le_def proof (rule conjI, fastforce)+ have daM: "adjoint M \ carrier_mat n n" using dM by auto have dBmA: "B - A \ carrier_mat n n" using dB dA by fastforce have "positive (B - A)" using le lowner_le_def by auto then have p: "positive (adjoint M * (B - A) * M)" using positive_close_under_left_right_mult_adjoint[OF daM dBmA] adjoint_adjoint[of M] by auto moreover have e: "adjoint M * (B - A) * M = adjoint M * B * M - adjoint M * A * M" using dM dB dA by (mat_assoc n) ultimately show "positive (adjoint M * B * M - adjoint M * A * M)" by auto qed lemma smult_distrib_left_minus_mat: fixes A B :: "'a::comm_ring_1 mat" assumes "A \ carrier_mat n n" "B \ carrier_mat n n" shows "c \\<^sub>m (B - A) = c \\<^sub>m B - c \\<^sub>m A" using assms by (auto simp add: minus_add_uminus_mat add_smult_distrib_left_mat) lemma lowner_le_smultc: fixes c :: complex assumes "c \ 0" "A \\<^sub>L B" "A \ carrier_mat n n" "B \ carrier_mat n n" shows "c \\<^sub>m A \\<^sub>L c \\<^sub>m B" proof - have eqBA: "c \\<^sub>m (B - A) = c \\<^sub>m B - c \\<^sub>m A" using assms by (auto simp add: smult_distrib_left_minus_mat) have "positive (B - A)" using assms(2) unfolding lowner_le_def by auto then have "positive (c \\<^sub>m (B - A))" using positive_smult[of "B-A" n c] assms by fastforce moreover have "c \\<^sub>m A \ carrier_mat n n" using index_smult_mat(2,3) assms(3) by auto moreover have "c \\<^sub>m B \ carrier_mat n n" using index_smult_mat(2,3) assms(4) by auto ultimately show ?thesis unfolding lowner_le_def using eqBA by fastforce qed lemma lowner_le_smult: fixes c :: real assumes "c \ 0" "A \\<^sub>L B" "A \ carrier_mat n n" "B \ carrier_mat n n" shows "c \\<^sub>m A \\<^sub>L c \\<^sub>m B" apply (rule lowner_le_smultc) using assms by (auto simp: less_eq_complex_def) lemma minus_smult_vec_distrib: fixes w :: "'a::comm_ring_1 vec" shows "(a - b) \\<^sub>v w = a \\<^sub>v w - b \\<^sub>v w" apply (rule eq_vecI) by (auto simp add: scalar_prod_def algebra_simps) lemma smult_mat_mult_mat_vec_assoc: fixes A :: "'a::comm_ring_1 mat" assumes A: "A \ carrier_mat n m" and w: "w \ carrier_vec m" shows "a \\<^sub>m A *\<^sub>v w = a \\<^sub>v (A *\<^sub>v w)" apply (rule eq_vecI) apply (simp add: scalar_prod_def carrier_matD[OF A] carrier_vecD[OF w]) apply (subst sum_distrib_left) apply (rule sum.cong, simp) by auto lemma mult_mat_vec_smult_vec_assoc: fixes A :: "'a::comm_ring_1 mat" assumes A: "A \ carrier_mat n m" and w: "w \ carrier_vec m" shows "A *\<^sub>v (a \\<^sub>v w) = a \\<^sub>v (A *\<^sub>v w)" apply (rule eq_vecI) apply (simp add: scalar_prod_def carrier_matD[OF A] carrier_vecD[OF w]) apply (subst sum_distrib_left) apply (rule sum.cong, simp) by auto lemma outer_prod_left_right_mat: fixes A B :: "complex mat" assumes du: "u \ carrier_vec d2" and dv: "v \ carrier_vec d3" and dA: "A \ carrier_mat d1 d2" and dB: "B \ carrier_mat d3 d4" shows "A * (outer_prod u v) * B = (outer_prod (A *\<^sub>v u) (adjoint B *\<^sub>v v))" unfolding outer_prod_def proof - have eq1: "A * (mat (dim_vec u) 1 (\(i, j). u $ i)) = mat (dim_vec (A *\<^sub>v u)) 1 (\(i, j). (A *\<^sub>v u) $ i)" apply (rule eq_matI) by (auto simp add: dA du scalar_prod_def) have conj: "conjugate a * b = conjugate ((a::complex) * conjugate b) " for a b by auto have eq2: "mat 1 (dim_vec v) (\(i, y). conjugate v $ y) * B = mat 1 (dim_vec (adjoint B *\<^sub>v v)) (\(i, y). conjugate (adjoint B *\<^sub>v v) $ y)" apply (rule eq_matI) apply (auto simp add: carrier_matD[OF dB] carrier_vecD[OF dv] scalar_prod_def adjoint_def conjugate_vec_def sum_conjugate ) apply (rule sum.cong) by (auto simp add: conj) have "A * (mat (dim_vec u) 1 (\(i, j). u $ i) * mat 1 (dim_vec v) (\(i, y). conjugate v $ y)) * B = (A * (mat (dim_vec u) 1 (\(i, j). u $ i))) *(mat 1 (dim_vec v) (\(i, y). conjugate v $ y)) * B" using dA du dv dB assoc_mult_mat[OF dA, of "mat (dim_vec u) 1 (\(i, j). u $ i)" 1 "mat 1 (dim_vec v) (\(i, y). conjugate v $ y)"] by fastforce also have "\ = (A * (mat (dim_vec u) 1 (\(i, j). u $ i))) *((mat 1 (dim_vec v) (\(i, y). conjugate v $ y)) * B)" using dA du dv dB assoc_mult_mat[OF _ _ dB, of "(A * (mat (dim_vec u) 1 (\(i, j). u $ i)))" d1 1] by fastforce finally show "A * (mat (dim_vec u) 1 (\(i, j). u $ i) * mat 1 (dim_vec v) (\(i, y). conjugate v $ y)) * B = mat (dim_vec (A *\<^sub>v u)) 1 (\(i, j). (A *\<^sub>v u) $ i) * mat 1 (dim_vec (adjoint B *\<^sub>v v)) (\(i, y). conjugate (adjoint B *\<^sub>v v) $ y)" using eq1 eq2 by auto qed subsection \Density operators\ definition density_operator :: "complex mat \ bool" where "density_operator A \ positive A \ trace A = 1" definition partial_density_operator :: "complex mat \ bool" where "partial_density_operator A \ positive A \ trace A \ 1" lemma pure_state_self_outer_prod_is_partial_density_operator: fixes v :: "complex vec" assumes dimv: "v \ carrier_vec n" and nv: "vec_norm v = 1" shows "partial_density_operator (outer_prod v v)" unfolding partial_density_operator_def proof have dimov: "outer_prod v v \ carrier_mat n n" using dimv by auto show "positive (outer_prod v v)" unfolding positive_def proof (rule, simp add: carrier_matD(2)[OF dimov] dimov, rule allI, rule impI) fix w assume "dim_vec (w::complex vec) = dim_col (outer_prod v v)" then have dimw: "w \ carrier_vec n" using dimov carrier_vecI by auto then have "inner_prod w ((outer_prod v v) *\<^sub>v w) = inner_prod w v * inner_prod v w" using inner_prod_outer_prod dimw dimv by auto also have "\ = inner_prod w v * conjugate (inner_prod w v)" using dimw dimv apply (subst conjugate_scalar_prod[of v "conjugate w"], simp) apply (subst conjugate_vec_sprod_comm[of "conjugate v" _ "conjugate w"], auto) apply (rule carrier_vec_conjugate[OF dimv]) apply (rule carrier_vec_conjugate[OF dimw]) done also have "\ \ 0" by (auto simp: less_eq_complex_def) finally show "inner_prod w ((outer_prod v v) *\<^sub>v w) \ 0". qed have eq: "trace (outer_prod v v) = (\i=0..i=0..i=0.. 1" by auto qed (* Lemma 2.1 *) lemma lowner_le_trace: assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" shows "A \\<^sub>L B \ (\\\carrier_mat n n. partial_density_operator \ \ trace (A * \) \ trace (B * \))" proof (rule iffI) have dimBmA: "B - A \ carrier_mat n n" using A B by auto { assume "A \\<^sub>L B" then have pBmA: "positive (B - A)" using lowner_le_def by auto moreover have "B - A \ carrier_mat n n" using assms by auto ultimately have "\M\carrier_mat n n. M * adjoint M = B - A" using positive_iff_decomp[of "B - A"] by auto then obtain M where dimM: "M \ carrier_mat n n" and M: "M * adjoint M = B - A" by auto { fix \ assume dimr: "\ \ carrier_mat n n" and pdr: "partial_density_operator \" have eq: "trace(B * \) - trace(A * \) = trace((B - A) * \)" using A B dimr apply (subst minus_mult_distrib_mat, auto) apply (subst trace_minus_linear, auto) done have pr: "positive \" using pdr partial_density_operator_def by auto then have "\P\carrier_mat n n. \ = P * adjoint P" using positive_iff_decomp dimr by auto then obtain P where dimP: "P \ carrier_mat n n" and P: "\ = P * adjoint P" by auto have "trace((B - A) * \) = trace(M * adjoint M * (P * adjoint P))" using P M by auto also have "\ = trace((adjoint P * M) * adjoint (adjoint P * M))" using dimM dimP by (mat_assoc n) also have "\ \ 0" using trace_adjoint_positive by auto finally have "trace((B - A) * \) \ 0". with eq have " trace (B * \) - trace (A * \) \ 0" by auto } then show "\\\carrier_mat n n. partial_density_operator \ \ trace (A * \) \ trace (B * \)" by auto } { assume asm: "\\\carrier_mat n n. partial_density_operator \ \ trace (A * \) \ trace (B * \)" have "positive (B - A)" proof - { fix v assume "dim_vec (v::complex vec) = dim_col (B - A) \ vec_norm v = 1" then have dimv: "v \ carrier_vec n" and nv: "vec_norm v = 1" using carrier_matD[OF dimBmA] by (auto intro: carrier_vecI) have dimov: "outer_prod v v \ carrier_mat n n" using dimv by auto then have "partial_density_operator (outer_prod v v)" using dimv nv pure_state_self_outer_prod_is_partial_density_operator by auto then have leq: "trace(A * (outer_prod v v)) \ trace(B * (outer_prod v v))" using asm dimov by auto have "trace((B - A) * (outer_prod v v)) = trace(B * (outer_prod v v)) - trace(A * (outer_prod v v))" using A B dimov apply (subst minus_mult_distrib_mat, auto) apply (subst trace_minus_linear, auto) done then have "trace((B - A) * (outer_prod v v)) \ 0" using leq by auto then have "inner_prod v ((B - A) *\<^sub>v v) \ 0" using trace_outer_prod_right[OF dimBmA dimv dimv] by auto } then show "positive (B - A)" using positive_iff_normalized_vec[of "B - A"] dimBmA A by simp qed then show "A \\<^sub>L B" using lowner_le_def A B by auto } qed lemma lowner_le_traceI: assumes "A \ carrier_mat n n" and "B \ carrier_mat n n" and "\\. \ \ carrier_mat n n \ partial_density_operator \ \ trace (A * \) \ trace (B * \)" shows "A \\<^sub>L B" using lowner_le_trace assms by auto lemma trace_pdo_eq_imp_eq: assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and teq: "\\. \ \ carrier_mat n n \ partial_density_operator \ \ trace (A * \) = trace (B * \)" shows "A = B" proof - from teq have "A \\<^sub>L B" using lowner_le_trace[OF A B] teq by auto moreover from teq have "B \\<^sub>L A" using lowner_le_trace[OF B A] teq by auto ultimately show "A = B" using lowner_le_antisym A B by auto qed lemma lowner_le_traceD: assumes "A \ carrier_mat n n" "B \ carrier_mat n n" "\ \ carrier_mat n n" and "A \\<^sub>L B" and "partial_density_operator \" shows "trace (A * \) \ trace (B * \)" using lowner_le_trace assms by blast lemma sum_only_one_neq_0: assumes "finite A" and "j \ A" and "\i. i \ A \ i \ j \ g i = 0" shows "sum g A = g j" proof - have "{j} \ A" using assms by auto moreover have "\i\A - {j}. g i = 0" using assms by simp ultimately have "sum g A = sum g {j}" using assms by (auto simp add: comm_monoid_add_class.sum.mono_neutral_right[of A "{j}" g]) moreover have "sum g {j} = g j" by simp ultimately show ?thesis by auto qed end diff --git a/thys/Saturation_Framework/Given_Clause_Architectures.thy b/thys/Saturation_Framework/Given_Clause_Architectures.thy --- a/thys/Saturation_Framework/Given_Clause_Architectures.thy +++ b/thys/Saturation_Framework/Given_Clause_Architectures.thy @@ -1,1174 +1,1174 @@ (* Title: Given Clause Prover Architectures * Author: Sophie Tourret , 2019-2020 *) section \Given Clause Prover Architectures\ text \This section covers all the results presented in the section 4 of the report. This is where abstract architectures of provers are defined and proven dynamically refutationally complete.\ theory Given_Clause_Architectures imports Lambda_Free_RPOs.Lambda_Free_Util Labeled_Lifting_to_Non_Ground_Calculi begin subsection \Basis of the Given Clause Prover Architectures\ locale given_clause_basis = std?: labeled_lifting_intersection Bot_F Inf_F Bot_G Q entails_q Inf_G_q Red_I_q Red_F_q \_F_q \_I_q Inf_FL for Bot_F :: "'f set" and Inf_F :: "'f inference set" and Bot_G :: "'g set" and Q :: "'q set" and entails_q :: "'q \ 'g set \ 'g set \ bool" and Inf_G_q :: \'q \ 'g inference set\ and Red_I_q :: "'q \ 'g set \ 'g inference set" and Red_F_q :: "'q \ 'g set \ 'g set" and \_F_q :: "'q \ 'f \ 'g set" and \_I_q :: "'q \ 'f inference \ 'g inference set option" and Inf_FL :: \('f \ 'l) inference set\ + fixes Equiv_F :: "'f \ 'f \ bool" (infix "\" 50) and Prec_F :: "'f \ 'f \ bool" (infix "\\" 50) and Prec_L :: "'l \ 'l \ bool" (infix "\L" 50) and active :: "'l" assumes equiv_equiv_F: "equivp (\)" and wf_prec_F: "minimal_element (\\) UNIV" and wf_prec_L: "minimal_element (\L) UNIV" and compat_equiv_prec: "C1 \ D1 \ C2 \ D2 \ C1 \\ C2 \ D1 \\ D2" and equiv_F_grounding: "q \ Q \ C1 \ C2 \ \_F_q q C1 \ \_F_q q C2" and prec_F_grounding: "q \ Q \ C2 \\ C1 \ \_F_q q C1 \ \_F_q q C2" and active_minimal: "l2 \ active \ active \L l2" and at_least_two_labels: "\l2. active \L l2" and inf_never_active: "\ \ Inf_FL \ snd (concl_of \) \ active" and static_ref_comp: "statically_complete_calculus Bot_F Inf_F (\\\) no_labels.Red_I_\ no_labels.Red_F_\_empty" begin abbreviation Prec_eq_F :: "'f \ 'f \ bool" (infix "\\" 50) where "C \\ D \ C \ D \ C \\ D" definition Prec_FL :: "('f \ 'l) \ ('f \ 'l) \ bool" (infix "\" 50) where "Cl1 \ Cl2 \ fst Cl1 \\ fst Cl2 \ (fst Cl1 \ fst Cl2 \ snd Cl1 \L snd Cl2)" lemma irrefl_prec_F: "\ C \\ C" by (simp add: minimal_element.po[OF wf_prec_F, unfolded po_on_def irreflp_on_def]) lemma trans_prec_F: "C1 \\ C2 \ C2 \\ C3 \ C1 \\ C3" by (auto intro: minimal_element.po[OF wf_prec_F, unfolded po_on_def transp_on_def, THEN conjunct2, simplified, rule_format]) lemma wf_prec_FL: "minimal_element (\) UNIV" proof show "po_on (\) UNIV" unfolding po_on_def proof show "irreflp_on (\) UNIV" unfolding irreflp_on_def Prec_FL_def proof fix Cl assume a_in: "Cl \ (UNIV::('f \ 'l) set)" have "\ (fst Cl \\ fst Cl)" using wf_prec_F minimal_element.min_elt_ex by force moreover have "\ (snd Cl \L snd Cl)" using wf_prec_L minimal_element.min_elt_ex by force ultimately show "\ (fst Cl \\ fst Cl \ fst Cl \ fst Cl \ snd Cl \L snd Cl)" by blast qed next show "transp_on (\) UNIV" unfolding transp_on_def Prec_FL_def proof (simp, intro allI impI) fix C1 l1 C2 l2 C3 l3 assume trans_hyp: "(C1 \\ C2 \ C1 \ C2 \ l1 \L l2) \ (C2 \\ C3 \ C2 \ C3 \ l2 \L l3)" have "C1 \\ C2 \ C2 \ C3 \ C1 \\ C3" using compat_equiv_prec by (metis equiv_equiv_F equivp_def) moreover have "C1 \ C2 \ C2 \\ C3 \ C1 \\ C3" using compat_equiv_prec by (metis equiv_equiv_F equivp_def) moreover have "l1 \L l2 \ l2 \L l3 \ l1 \L l3" using wf_prec_L unfolding minimal_element_def po_on_def transp_on_def by (meson UNIV_I) moreover have "C1 \ C2 \ C2 \ C3 \ C1 \ C3" using equiv_equiv_F by (meson equivp_transp) ultimately show "C1 \\ C3 \ C1 \ C3 \ l1 \L l3" using trans_hyp using trans_prec_F by blast qed qed next show "wfp_on (\) UNIV" unfolding wfp_on_def proof assume contra: "\f. \i. f i \ UNIV \ f (Suc i) \ f i" then obtain f where f_suc: "\i. f (Suc i) \ f i" by blast define R :: "(('f \ 'l) \ ('f \ 'l)) set" where "R = {(Cl1, Cl2). fst Cl1 \\ fst Cl2}" define S :: "(('f \ 'l) \ ('f \ 'l)) set" where "S = {(Cl1, Cl2). fst Cl1 \ fst Cl2 \ snd Cl1 \L snd Cl2}" obtain k where f_chain: "\i. (f (Suc (i + k)), f (i + k)) \ S" proof (atomize_elim, rule wf_infinite_down_chain_compatible[of R f S]) show "wf R" unfolding R_def using wf_app[OF wf_prec_F[unfolded minimal_element_def, THEN conjunct2, unfolded wfp_on_UNIV wfP_def]] by force next show "\i. (f (Suc i), f i) \ R \ S" using f_suc unfolding R_def S_def Prec_FL_def by blast next show "R O S \ R" unfolding R_def S_def using compat_equiv_prec equiv_equiv_F equivp_reflp by fastforce qed define g where "\i. g i = f (i + k)" have g_chain: "\i. (g (Suc i), g i) \ S" unfolding g_def using f_chain by simp have wf_s: "wf S" unfolding S_def by (rule wf_subset[OF wf_app[OF wf_prec_L[unfolded minimal_element_def, THEN conjunct2, unfolded wfp_on_UNIV wfP_def], of snd]]) fast show False using g_chain[unfolded S_def] wf_s[unfolded S_def, folded wfP_def wfp_on_UNIV, unfolded wfp_on_def] by auto qed qed definition active_subset :: "('f \ 'l) set \ ('f \ 'l) set" where "active_subset M = {CL \ M. snd CL = active}" definition passive_subset :: "('f \ 'l) set \ ('f \ 'l) set" where "passive_subset M = {CL \ M. snd CL \ active}" lemma active_subset_insert[simp]: "active_subset (insert Cl N) = (if snd Cl = active then {Cl} else {}) \ active_subset N" unfolding active_subset_def by auto lemma active_subset_union[simp]: "active_subset (M \ N) = active_subset M \ active_subset N" unfolding active_subset_def by auto lemma passive_subset_insert[simp]: "passive_subset (insert Cl N) = (if snd Cl \ active then {Cl} else {}) \ passive_subset N" unfolding passive_subset_def by auto lemma passive_subset_union[simp]: "passive_subset (M \ N) = passive_subset M \ passive_subset N" unfolding passive_subset_def by auto sublocale std?: statically_complete_calculus Bot_FL Inf_FL "(\\\L)" Red_I Red_F using labeled_static_ref[OF static_ref_comp] . lemma labeled_tiebreaker_lifting: assumes q_in: "q \ Q" shows "tiebreaker_lifting Bot_FL Inf_FL Bot_G (entails_q q) (Inf_G_q q) (Red_I_q q) (Red_F_q q) (\_F_L_q q) (\_I_L_q q) (\g. Prec_FL)" proof - have "tiebreaker_lifting Bot_FL Inf_FL Bot_G (entails_q q) (Inf_G_q q) (Red_I_q q) (Red_F_q q) (\_F_L_q q) (\_I_L_q q) (\g Cl Cl'. False)" using ord_fam_lifted_q[OF q_in] . then have "standard_lifting Inf_FL Bot_G (Inf_G_q q) (entails_q q) (Red_I_q q) (Red_F_q q) Bot_FL (\_F_L_q q) (\_I_L_q q)" using lifted_q[OF q_in] by blast then show "tiebreaker_lifting Bot_FL Inf_FL Bot_G (entails_q q) (Inf_G_q q) (Red_I_q q) (Red_F_q q) (\_F_L_q q) (\_I_L_q q) (\g. Prec_FL)" using wf_prec_FL by (simp add: tiebreaker_lifting.intro tiebreaker_lifting_axioms.intro) qed sublocale lifting_intersection Inf_FL Bot_G Q Inf_G_q entails_q Red_I_q Red_F_q Bot_FL \_F_L_q \_I_L_q "\g. Prec_FL" using labeled_tiebreaker_lifting unfolding lifting_intersection_def by (simp add: lifting_intersection_axioms.intro no_labels.ground.consequence_relation_family_axioms no_labels.ground.inference_system_family_axioms) notation derive (infix "\L" 50) lemma std_Red_I_eq: "std.Red_I = Red_I_\" unfolding Red_I_\_q_def Red_I_\_L_q_def by simp lemma std_Red_F_eq: "std.Red_F = Red_F_\_empty" unfolding Red_F_\_empty_q_def Red_F_\_empty_L_q_def by simp sublocale statically_complete_calculus Bot_FL Inf_FL "(\\\L)" Red_I Red_F by unfold_locales (use statically_complete std_Red_I_eq in auto) (* lem:redundant-labeled-inferences *) lemma labeled_red_inf_eq_red_inf: assumes i_in: "\ \ Inf_FL" shows "\ \ Red_I N \ to_F \ \ no_labels.Red_I_\ (fst ` N)" proof assume i_in2: "\ \ Red_I N" then have "X \ Red_I_\_q ` Q \ \ \ X N" for X unfolding Red_I_def by blast obtain X0 where "X0 \ Red_I_\_q ` Q" using Q_nonempty by blast then obtain q0 where x0_is: "X0 N = Red_I_\_q q0 N" by blast then obtain Y0 where y0_is: "Y0 (fst ` N) = to_F ` (X0 N)" by auto have "Y0 (fst ` N) = no_labels.Red_I_\_q q0 (fst ` N)" unfolding y0_is proof show "to_F ` X0 N \ no_labels.Red_I_\_q q0 (fst ` N)" proof fix \0 assume i0_in: "\0 \ to_F ` X0 N" then have i0_in2: "\0 \ to_F ` Red_I_\_q q0 N" using x0_is by argo then obtain \0_FL where i0_FL_in: "\0_FL \ Inf_FL" and i0_to_i0_FL: "\0 = to_F \0_FL" and subs1: "((\_I_L_q q0 \0_FL) \ None \ the (\_I_L_q q0 \0_FL) \ Red_I_q q0 (\_Fset_q q0 N)) \ ((\_I_L_q q0 \0_FL = None) \ \_F_L_q q0 (concl_of \0_FL) \ \_Fset_q q0 N \ Red_F_q q0 (\_Fset_q q0 N))" unfolding Red_I_\_q_def by blast have concl_swap: "fst (concl_of \0_FL) = concl_of \0" unfolding concl_of_def i0_to_i0_FL to_F_def by simp have i0_in3: "\0 \ Inf_F" using i0_to_i0_FL Inf_FL_to_Inf_F[OF i0_FL_in] unfolding to_F_def by blast { assume not_none: "\_I_q q0 \0 \ None" and "the (\_I_q q0 \0) \ {}" then obtain \1 where i1_in: "\1 \ the (\_I_q q0 \0)" by blast have "the (\_I_q q0 \0) \ Red_I_q q0 (no_labels.\_Fset_q q0 (fst ` N))" using subs1 i0_to_i0_FL not_none by auto } moreover { assume is_none: "\_I_q q0 \0 = None" then have "\_F_q q0 (concl_of \0) \ no_labels.\_Fset_q q0 (fst ` N) \ Red_F_q q0 (no_labels.\_Fset_q q0 (fst ` N))" using subs1 i0_to_i0_FL concl_swap by simp } ultimately show "\0 \ no_labels.Red_I_\_q q0 (fst ` N)" unfolding no_labels.Red_I_\_q_def using i0_in3 by auto qed next show "no_labels.Red_I_\_q q0 (fst ` N) \ to_F ` X0 N" proof fix \0 assume i0_in: "\0 \ no_labels.Red_I_\_q q0 (fst ` N)" then have i0_in2: "\0 \ Inf_F" unfolding no_labels.Red_I_\_q_def by blast obtain \0_FL where i0_FL_in: "\0_FL \ Inf_FL" and i0_to_i0_FL: "\0 = to_F \0_FL" using Inf_F_to_Inf_FL[OF i0_in2] unfolding to_F_def by (metis Ex_list_of_length fst_conv inference.exhaust_sel inference.inject map_fst_zip) have concl_swap: "fst (concl_of \0_FL) = concl_of \0" unfolding concl_of_def i0_to_i0_FL to_F_def by simp have subs1: "((\_I_L_q q0 \0_FL) \ None \ the (\_I_L_q q0 \0_FL) \ Red_I_q q0 (\_Fset_q q0 N)) \ ((\_I_L_q q0 \0_FL = None) \ \_F_L_q q0 (concl_of \0_FL) \ (\_Fset_q q0 N \ Red_F_q q0 (\_Fset_q q0 N)))" using i0_in i0_to_i0_FL concl_swap unfolding no_labels.Red_I_\_q_def by simp then have "\0_FL \ Red_I_\_q q0 N" using i0_FL_in unfolding Red_I_\_q_def by simp then show "\0 \ to_F ` X0 N" using x0_is i0_to_i0_FL i0_in2 by blast qed qed then have "Y \ no_labels.Red_I_\_q ` Q \ to_F \ \ Y (fst ` N)" for Y using i_in2 no_labels.Red_I_def std_Red_I_eq red_inf_impl by force then show "to_F \ \ no_labels.Red_I_\ (fst ` N)" unfolding Red_I_def no_labels.Red_I_\_def by blast next assume to_F_in: "to_F \ \ no_labels.Red_I_\ (fst ` N)" have imp_to_F: "X \ no_labels.Red_I_\_q ` Q \ to_F \ \ X (fst ` N)" for X using to_F_in unfolding no_labels.Red_I_\_def by blast then have to_F_in2: "to_F \ \ no_labels.Red_I_\_q q (fst ` N)" if "q \ Q" for q using that by auto have "Red_I_\_q q N = {\0_FL \ Inf_FL. to_F \0_FL \ no_labels.Red_I_\_q q (fst ` N)}" for q proof show "Red_I_\_q q N \ {\0_FL \ Inf_FL. to_F \0_FL \ no_labels.Red_I_\_q q (fst ` N)}" proof fix q0 \1 assume i1_in: "\1 \ Red_I_\_q q0 N" have i1_in2: "\1 \ Inf_FL" using i1_in unfolding Red_I_\_q_def by blast then have to_F_i1_in: "to_F \1 \ Inf_F" using Inf_FL_to_Inf_F unfolding to_F_def by simp have concl_swap: "fst (concl_of \1) = concl_of (to_F \1)" unfolding concl_of_def to_F_def by simp then have i1_to_F_in: "to_F \1 \ no_labels.Red_I_\_q q0 (fst ` N)" using i1_in to_F_i1_in unfolding Red_I_\_q_def no_labels.Red_I_\_q_def by force show "\1 \ {\0_FL \ Inf_FL. to_F \0_FL \ no_labels.Red_I_\_q q0 (fst ` N)}" using i1_in2 i1_to_F_in by blast qed next show "{\0_FL \ Inf_FL. to_F \0_FL \ no_labels.Red_I_\_q q (fst ` N)} \ Red_I_\_q q N" proof fix q0 \1 assume i1_in: "\1 \ {\0_FL \ Inf_FL. to_F \0_FL \ no_labels.Red_I_\_q q0 (fst ` N)}" then have i1_in2: "\1 \ Inf_FL" by blast then have to_F_i1_in: "to_F \1 \ Inf_F" using Inf_FL_to_Inf_F unfolding to_F_def by simp have concl_swap: "fst (concl_of \1) = concl_of (to_F \1)" unfolding concl_of_def to_F_def by simp then have "((\_I_L_q q0 \1) \ None \ the (\_I_L_q q0 \1) \ Red_I_q q0 (\_Fset_q q0 N)) \ (\_I_L_q q0 \1 = None \ \_F_L_q q0 (concl_of \1) \ \_Fset_q q0 N \ Red_F_q q0 (\_Fset_q q0 N))" using i1_in unfolding no_labels.Red_I_\_q_def by auto then show "\1 \ Red_I_\_q q0 N" using i1_in2 unfolding Red_I_\_q_def by blast qed qed then have "\ \ Red_I_\_q q N" if "q \ Q" for q using that to_F_in2 i_in unfolding Red_I_\_q_def no_labels.Red_I_\_q_def by auto then show "\ \ Red_I_\ N" unfolding Red_I_\_def by blast qed (* lem:redundant-labeled-formulas *) lemma red_labeled_clauses: assumes \C \ no_labels.Red_F_\_empty (fst ` N) \ (\C' \ fst ` N. C' \\ C) \ (\(C', L') \ N. L' \L L \ C' \\ C)\ shows \(C, L) \ Red_F N\ proof - note assms moreover have i: \C \ no_labels.Red_F_\_empty (fst ` N) \ (C, L) \ Red_F N\ proof - assume "C \ no_labels.Red_F_\_empty (fst ` N)" then have "C \ no_labels.Red_F_\_empty_q q (fst ` N)" if "q \ Q" for q unfolding no_labels.Red_F_\_empty_def using that by fast then have g_in_red: "\_F_q q C \ Red_F_q q (no_labels.\_Fset_q q (fst ` N))" if "q \ Q" for q unfolding no_labels.Red_F_\_empty_q_def using that by blast have "\_F_L_q q (C, L) \ Red_F_q q (\_Fset_q q N)" if "q \ Q" for q using that g_in_red by simp then show ?thesis unfolding Red_F_def Red_F_\_q_def by blast qed moreover have ii: \\C' \ fst ` N. C' \\ C \ (C, L) \ Red_F N\ proof - assume "\C' \ fst ` N. C' \\ C" then obtain C' where c'_in: "C' \ fst ` N" and c_prec_c': "C' \\ C" by blast obtain L' where c'_l'_in: "(C', L') \ N" using c'_in by auto have c'_l'_prec: "(C', L') \ (C, L)" using c_prec_c' unfolding Prec_FL_def by simp have c_in_c'_g: "\_F_q q C \ \_F_q q C'" if "q \ Q" for q using prec_F_grounding[OF that c_prec_c'] by presburger then have "\_F_L_q q (C, L) \ \_F_L_q q (C', L')" if "q \ Q" for q using that by auto then have "(C, L) \ Red_F_\_q q N" if "q \ Q" for q unfolding Red_F_\_q_def using that c'_l'_in c'_l'_prec by blast then show ?thesis unfolding Red_F_def by blast qed moreover have iii: \\(C', L') \ N. L' \L L \ C' \\ C \ (C, L) \ Red_F N\ proof - assume "\(C', L') \ N. L' \L L \ C' \\ C" then obtain C' L' where c'_l'_in: "(C', L') \ N" and l'_sub_l: "L' \L L" and c'_sub_c: "C' \\ C" by fast have "(C, L) \ Red_F N" if "C' \\ C" using that c'_l'_in ii by fastforce moreover { assume equiv_c_c': "C \ C'" then have equiv_c'_c: "C' \ C" using equiv_equiv_F by (simp add: equivp_symp) then have c'_l'_prec: "(C', L') \ (C, L)" using l'_sub_l unfolding Prec_FL_def by simp have "\_F_q q C = \_F_q q C'" if "q \ Q" for q using that equiv_F_grounding equiv_c_c' equiv_c'_c by (simp add: set_eq_subset) then have "\_F_L_q q (C, L) = \_F_L_q q (C', L')" if "q \ Q" for q using that by auto then have "(C, L) \ Red_F_\_q q N" if "q \ Q" for q unfolding Red_F_\_q_def using that c'_l'_in c'_l'_prec by blast then have ?thesis unfolding Red_F_def by blast } ultimately show ?thesis using c'_sub_c equiv_equiv_F equivp_symp by fastforce qed ultimately show ?thesis by blast qed end subsection \Given Clause Procedure\ locale given_clause = given_clause_basis Bot_F Inf_F Bot_G Q entails_q Inf_G_q Red_I_q Red_F_q \_F_q \_I_q Inf_FL Equiv_F Prec_F Prec_L active for Bot_F :: "'f set" and Inf_F :: "'f inference set" and Bot_G :: "'g set" and Q :: "'q set" and entails_q :: "'q \ 'g set \ 'g set \ bool" and Inf_G_q :: \'q \ 'g inference set\ and Red_I_q :: "'q \ 'g set \ 'g inference set" and Red_F_q :: "'q \ 'g set \ 'g set" and \_F_q :: "'q \ 'f \ 'g set" and \_I_q :: "'q \ 'f inference \ 'g inference set option" and Inf_FL :: \('f \ 'l) inference set\ and Equiv_F :: "'f \ 'f \ bool" (infix "\" 50) and Prec_F :: "'f \ 'f \ bool" (infix "\\" 50) and Prec_L :: "'l \ 'l \ bool" (infix "\L" 50) and active :: 'l + assumes inf_have_prems: "\F \ Inf_F \ prems_of \F \ []" begin lemma labeled_inf_have_prems: "\ \ Inf_FL \ prems_of \ \ []" using inf_have_prems Inf_FL_to_Inf_F by fastforce inductive step :: "('f \ 'l) set \ ('f \ 'l) set \ bool" (infix "\GC" 50) where process: "N1 = N \ M \ N2 = N \ M' \ M \ Red_F (N \ M') \ active_subset M' = {} \ N1 \GC N2" | infer: "N1 = N \ {(C, L)} \ N2 = N \ {(C, active)} \ M \ L \ active \ active_subset M = {} \ no_labels.Inf_between (fst ` (active_subset N)) {C} \ no_labels.Red_I (fst ` (N \ {(C, active)} \ M)) \ N1 \GC N2" lemma one_step_equiv: "N1 \GC N2 \ N1 \L N2" proof (cases N1 N2 rule: step.cases) show "N1 \GC N2 \ N1 \GC N2" by blast next fix N M M' assume gc_step: "N1 \GC N2" and n1_is: "N1 = N \ M" and n2_is: "N2 = N \ M'" and m_red: "M \ Red_F (N \ M')" and active_empty: "active_subset M' = {}" have "N1 - N2 \ Red_F N2" using n1_is n2_is m_red by auto then show "N1 \L N2" unfolding derive.simps by blast next fix N C L M assume gc_step: "N1 \GC N2" and n1_is: "N1 = N \ {(C, L)}" and not_active: "L \ active" and n2_is: "N2 = N \ {(C, active)} \ M" and active_empty: "active_subset M = {}" have "(C, active) \ N2" using n2_is by auto moreover have "C \\ C" using equiv_equiv_F by (metis equivp_def) moreover have "active \L L" using active_minimal[OF not_active] . ultimately have "{(C, L)} \ Red_F N2" using red_labeled_clauses by blast moreover have "N1 - N2 = {} \ N1 - N2 = {(C, L)}" using n1_is n2_is by blast ultimately have "N1 - N2 \ Red_F N2" using std_Red_F_eq by blast then show "N1 \L N2" unfolding derive.simps by blast qed (* lem:gc-derivations-are-red-derivations *) lemma gc_to_red: "chain (\GC) Ns \ chain (\L) Ns" using one_step_equiv Lazy_List_Chain.chain_mono by blast lemma (in-) all_ex_finite_set: "(\(j::nat)\{0..(n::nat). P j n) \ (\n1 n2. \j\{0.. P j n2 \ n1 = n2) \ finite {n. \j \ {0.. nat \ bool" assume allj_exn: "\j\{0..n. P j n" and uniq_n: "\n1 n2. \j\{0.. P j n2 \ n1 = n2" have "{n. \j \ {0..((\j. {n. P j n}) ` {0..j\{0.. finite {n. \j \ {0..j. {n. P j n}"] by simp have "\j\{0..!n. P j n" using allj_exn uniq_n by blast then have "\j\{0..j \ {0..GC) Ns" and init_state: "active_subset (lhd Ns) = {}" and final_state: "passive_subset (Liminf_llist Ns) = {}" shows "fair Ns" unfolding fair_def proof fix \ assume i_in: "\ \ Inf_from (Liminf_llist Ns)" note lhd_is = lhd_conv_lnth[OF chain_not_lnull[OF deriv]] have i_in_inf_fl: "\ \ Inf_FL" using i_in unfolding Inf_from_def by blast have "Liminf_llist Ns = active_subset (Liminf_llist Ns)" using final_state unfolding passive_subset_def active_subset_def by blast then have i_in2: "\ \ Inf_from (active_subset (Liminf_llist Ns))" using i_in by simp define m where "m = length (prems_of \)" then have m_def_F: "m = length (prems_of (to_F \))" unfolding to_F_def by simp have i_in_F: "to_F \ \ Inf_F" using i_in Inf_FL_to_Inf_F unfolding Inf_from_def to_F_def by blast then have m_pos: "m > 0" using m_def_F using inf_have_prems by blast have exist_nj: "\j \ {0..nj. enat (Suc nj) < llength Ns \ prems_of \ ! j \ active_subset (lnth Ns nj) \ (\k. k > nj \ enat k < llength Ns \ prems_of \ ! j \ active_subset (lnth Ns k)))" proof clarify fix j assume j_in: "j \ {0.. ! j" using i_in2 unfolding m_def Inf_from_def active_subset_def by (smt Collect_mem_eq Collect_mono_iff atLeastLessThan_iff nth_mem old.prod.exhaust snd_conv) then have "(C, active) \ Liminf_llist Ns" using j_in i_in unfolding m_def Inf_from_def by force then obtain nj where nj_is: "enat nj < llength Ns" and c_in2: "(C, active) \ \ (lnth Ns ` {k. nj \ k \ enat k < llength Ns})" unfolding Liminf_llist_def using init_state by blast then have c_in3: "\k. k \ nj \ enat k < llength Ns \ (C, active) \ lnth Ns k" by blast have nj_pos: "nj > 0" using init_state c_in2 nj_is unfolding active_subset_def lhd_is by force obtain nj_min where nj_min_is: "nj_min = (LEAST nj. enat nj < llength Ns \ (C, active) \ \ (lnth Ns ` {k. nj \ k \ enat k < llength Ns}))" by blast then have in_allk: "\k. k \ nj_min \ enat k < llength Ns \ (C, active) \ (lnth Ns k)" using c_in3 nj_is c_in2 by (metis (mono_tags, lifting) INT_E LeastI_ex mem_Collect_eq) have njm_smaller_D: "enat nj_min < llength Ns" using nj_min_is by (smt LeastI_ex \\thesis. (\nj. \enat nj < llength Ns; (C, active) \ \ (lnth Ns ` {k. nj \ k \ enat k < llength Ns})\ \ thesis) \ thesis\) have "nj_min > 0" using nj_is c_in2 nj_pos nj_min_is lhd_is by (metis (mono_tags, lifting) Collect_empty_eq \(C, active) \ Liminf_llist Ns\ \Liminf_llist Ns = active_subset (Liminf_llist Ns)\ \\k\nj_min. enat k < llength Ns \ (C, active) \ lnth Ns k\ active_subset_def init_state linorder_not_less mem_Collect_eq zero_enat_def chain_length_pos[OF deriv]) then obtain njm_prec where nj_prec_is: "Suc njm_prec = nj_min" using gr0_conv_Suc by auto then have njm_prec_njm: "njm_prec < nj_min" by blast then have njm_prec_njm_enat: "enat njm_prec < enat nj_min" by simp have njm_prec_smaller_d: "njm_prec < llength Ns" - using HOL.no_atp(15)[OF njm_smaller_D njm_prec_njm_enat] . + by (rule less_trans[OF njm_prec_njm_enat njm_smaller_D]) have njm_prec_all_suc: "\k>njm_prec. enat k < llength Ns \ (C, active) \ lnth Ns k" using nj_prec_is in_allk by simp have notin_njm_prec: "(C, active) \ lnth Ns njm_prec" proof (rule ccontr) assume "\ (C, active) \ lnth Ns njm_prec" then have absurd_hyp: "(C, active) \ lnth Ns njm_prec" by simp have prec_smaller: "enat njm_prec < llength Ns" using nj_min_is nj_prec_is by (smt LeastI_ex Suc_leD \\thesis. (\nj. \enat nj < llength Ns; (C, active) \ \ (lnth Ns ` {k. nj \ k \ enat k < llength Ns})\ \ thesis) \ thesis\ enat_ord_simps(1) le_eq_less_or_eq le_less_trans) have "(C, active) \ \ (lnth Ns ` {k. njm_prec \ k \ enat k < llength Ns})" proof - { fix k assume k_in: "njm_prec \ k \ enat k < llength Ns" have "k = njm_prec \ (C, active) \ lnth Ns k" using absurd_hyp by simp moreover have "njm_prec < k \ (C, active) \ lnth Ns k" using nj_prec_is in_allk k_in by simp ultimately have "(C, active) \ lnth Ns k" using k_in by fastforce } then show "(C, active) \ \ (lnth Ns ` {k. njm_prec \ k \ enat k < llength Ns})" by blast qed then have "enat njm_prec < llength Ns \ (C, active) \ \ (lnth Ns ` {k. njm_prec \ k \ enat k < llength Ns})" using prec_smaller by blast then show False using nj_min_is nj_prec_is Orderings.wellorder_class.not_less_Least njm_prec_njm by blast qed then have notin_active_subs_njm_prec: "(C, active) \ active_subset (lnth Ns njm_prec)" unfolding active_subset_def by blast then show "\nj. enat (Suc nj) < llength Ns \ prems_of \ ! j \ active_subset (lnth Ns nj) \ (\k. k > nj \ enat k < llength Ns \ prems_of \ ! j \ active_subset (lnth Ns k))" using c_is njm_prec_all_suc njm_prec_smaller_d by (metis (mono_tags, lifting) active_subset_def mem_Collect_eq nj_prec_is njm_smaller_D snd_conv) qed define nj_set where "nj_set = {nj. (\j\{0.. prems_of \ ! j \ active_subset (lnth Ns nj) \ (\k. k > nj \ enat k < llength Ns \ prems_of \ ! j \ active_subset (lnth Ns k)))}" then have nj_not_empty: "nj_set \ {}" proof - have zero_in: "0 \ {0.. ! 0 \ active_subset (lnth Ns n0)" and "\k>n0. enat k < llength Ns \ prems_of \ ! 0 \ active_subset (lnth Ns k)" using exist_nj by fast then have "n0 \ nj_set" unfolding nj_set_def using zero_in by blast then show "nj_set \ {}" by auto qed have nj_finite: "finite nj_set" using all_ex_finite_set[OF exist_nj] by (metis (no_types, lifting) Suc_ile_eq dual_order.strict_implies_order linorder_neqE_nat nj_set_def) (* the n below in the n-1 from the pen-and-paper proof *) have "\n \ nj_set. \nj \ nj_set. nj \ n" using nj_not_empty nj_finite using Max_ge Max_in by blast then obtain n where n_in: "n \ nj_set" and n_bigger: "\nj \ nj_set. nj \ n" by blast then obtain j0 where j0_in: "j0 \ {0.. ! j0 \ active_subset (lnth Ns n)" and j0_allin: "(\k. k > n \ enat k < llength Ns \ prems_of \ ! j0 \ active_subset (lnth Ns k))" unfolding nj_set_def by blast obtain C0 where C0_is: "prems_of \ ! j0 = (C0, active)" using j0_in using i_in2 unfolding m_def Inf_from_def active_subset_def by (smt Collect_mem_eq Collect_mono_iff atLeastLessThan_iff nth_mem old.prod.exhaust snd_conv) then have C0_prems_i: "(C0, active) \ set (prems_of \)" using in_set_conv_nth j0_in m_def by force have C0_in: "(C0, active) \ (lnth Ns (Suc n))" using C0_is j0_allin suc_n_length by (simp add: active_subset_def) have C0_notin: "(C0, active) \ (lnth Ns n)" using C0_is j0_notin unfolding active_subset_def by simp have step_n: "lnth Ns n \GC lnth Ns (Suc n)" using deriv chain_lnth_rel n_in unfolding nj_set_def by blast have "\N C L M. (lnth Ns n = N \ {(C, L)} \ lnth Ns (Suc n) = N \ {(C, active)} \ M \ L \ active \ active_subset M = {} \ no_labels.Inf_between (fst ` (active_subset N)) {C} \ no_labels.Red_I (fst ` (N \ {(C, active)} \ M)))" proof - have proc_or_infer: "(\N1 N M N2 M'. lnth Ns n = N1 \ lnth Ns (Suc n) = N2 \ N1 = N \ M \ N2 = N \ M' \ M \ Red_F (N \ M') \ active_subset M' = {}) \ (\N1 N C L N2 M. lnth Ns n = N1 \ lnth Ns (Suc n) = N2 \ N1 = N \ {(C, L)} \ N2 = N \ {(C, active)} \ M \ L \ active \ active_subset M = {} \ no_labels.Inf_between (fst ` (active_subset N)) {C} \ no_labels.Red_I (fst ` (N \ {(C, active)} \ M)))" using step.simps[of "lnth Ns n" "lnth Ns (Suc n)"] step_n by blast show ?thesis using C0_in C0_notin proc_or_infer j0_in C0_is by (smt Un_iff active_subset_def mem_Collect_eq snd_conv sup_bot.right_neutral) qed then obtain N M L where inf_from_subs: "no_labels.Inf_between (fst ` (active_subset N)) {C0} \ no_labels.Red_I (fst ` (N \ {(C0, active)} \ M))" and nth_d_is: "lnth Ns n = N \ {(C0, L)}" and suc_nth_d_is: "lnth Ns (Suc n) = N \ {(C0, active)} \ M" and l_not_active: "L \ active" using C0_in C0_notin j0_in C0_is using active_subset_def by fastforce have "j \ {0.. prems_of \ ! j \ prems_of \ ! j0 \ prems_of \ ! j \ (active_subset N)" for j proof - fix j assume j_in: "j \ {0.. ! j \ prems_of \ ! j0" obtain nj where nj_len: "enat (Suc nj) < llength Ns" and nj_prems: "prems_of \ ! j \ active_subset (lnth Ns nj)" and nj_greater: "(\k. k > nj \ enat k < llength Ns \ prems_of \ ! j \ active_subset (lnth Ns k))" using exist_nj j_in by blast then have "nj \ nj_set" unfolding nj_set_def using j_in by blast moreover have "nj \ n" proof (rule ccontr) assume "\ nj \ n" then have "prems_of \ ! j = (C0, active)" using C0_in C0_notin step.simps[of "lnth Ns n" "lnth Ns (Suc n)"] step_n by (smt Un_iff nth_d_is suc_nth_d_is l_not_active active_subset_def insertCI insertE lessI mem_Collect_eq nj_greater nj_prems snd_conv suc_n_length) then show False using j_not_j0 C0_is by simp qed ultimately have "nj < n" using n_bigger by force then have "prems_of \ ! j \ (active_subset (lnth Ns n))" using nj_greater n_in Suc_ile_eq dual_order.strict_implies_order unfolding nj_set_def by blast then show "prems_of \ ! j \ (active_subset N)" using nth_d_is l_not_active unfolding active_subset_def by force qed then have "set (prems_of \) \ active_subset N \ {(C0, active)}" using C0_prems_i C0_is m_def by (metis Un_iff atLeast0LessThan in_set_conv_nth insertCI lessThan_iff subrelI) moreover have "\ (set (prems_of \) \ active_subset N - {(C0, active)})" using C0_prems_i by blast ultimately have "\ \ Inf_between (active_subset N) {(C0, active)}" using i_in_inf_fl unfolding Inf_between_def Inf_from_def by blast then have "to_F \ \ no_labels.Inf_between (fst ` (active_subset N)) {C0}" unfolding to_F_def Inf_between_def Inf_from_def no_labels.Inf_between_def no_labels.Inf_from_def using Inf_FL_to_Inf_F by force then have "to_F \ \ no_labels.Red_I (fst ` (lnth Ns (Suc n)))" using suc_nth_d_is inf_from_subs by fastforce then have "\q \ Q. (\_I_q q (to_F \) \ None \ the (\_I_q q (to_F \)) \ Red_I_q q (\ (\_F_q q ` fst ` lnth Ns (Suc n)))) \ (\_I_q q (to_F \) = None \ \_F_q q (concl_of (to_F \)) \ \ (\_F_q q ` fst ` lnth Ns (Suc n)) \ Red_F_q q (\ (\_F_q q ` fst ` lnth Ns (Suc n))))" unfolding to_F_def no_labels.Red_I_def no_labels.Red_I_\_q_def by blast then have "\ \ Red_I_\ (lnth Ns (Suc n))" using i_in_inf_fl unfolding Red_I_\_def Red_I_\_q_def by (simp add: to_F_def) then show "\ \ Sup_llist (lmap Red_I_\ Ns)" unfolding Sup_llist_def using suc_n_length by auto qed theorem gc_complete_Liminf: assumes deriv: "chain (\GC) Ns" and init_state: "active_subset (lhd Ns) = {}" and final_state: "passive_subset (Liminf_llist Ns) = {}" and b_in: "B \ Bot_F" and bot_entailed: "no_labels.entails_\ (fst ` lhd Ns) {B}" shows "\BL \ Bot_FL. BL \ Liminf_llist Ns" proof - note lhd_is = lhd_conv_lnth[OF chain_not_lnull[OF deriv]] have labeled_b_in: "(B, active) \ Bot_FL" using b_in by simp have labeled_bot_entailed: "entails_\_L (lhd Ns) {(B, active)}" using labeled_entailment_lifting bot_entailed lhd_is by fastforce have fair: "fair Ns" using gc_fair[OF deriv init_state final_state] . then show ?thesis using dynamically_complete_Liminf[OF labeled_b_in gc_to_red[OF deriv] fair labeled_bot_entailed] by blast qed (* thm:gc-completeness *) theorem gc_complete: assumes deriv: "chain (\GC) Ns" and init_state: "active_subset (lhd Ns) = {}" and final_state: "passive_subset (Liminf_llist Ns) = {}" and b_in: "B \ Bot_F" and bot_entailed: "no_labels.entails_\ (fst ` lhd Ns) {B}" shows "\i. enat i < llength Ns \ (\BL \ Bot_FL. BL \ lnth Ns i)" proof - note lhd_is = lhd_conv_lnth[OF chain_not_lnull[OF deriv]] have "\BL\Bot_FL. BL \ Liminf_llist Ns" using assms by (rule gc_complete_Liminf) then show ?thesis unfolding Liminf_llist_def by auto qed end subsection \Lazy Given Clause Procedure\ locale lazy_given_clause = given_clause_basis Bot_F Inf_F Bot_G Q entails_q Inf_G_q Red_I_q Red_F_q \_F_q \_I_q Inf_FL Equiv_F Prec_F Prec_L active for Bot_F :: "'f set" and Inf_F :: "'f inference set" and Bot_G :: "'g set" and Q :: "'q set" and entails_q :: "'q \ 'g set \ 'g set \ bool" and Inf_G_q :: \'q \ 'g inference set\ and Red_I_q :: "'q \ 'g set \ 'g inference set" and Red_F_q :: "'q \ 'g set \ 'g set" and \_F_q :: "'q \ 'f \ 'g set" and \_I_q :: "'q \ 'f inference \ 'g inference set option" and Inf_FL :: \('f \ 'l) inference set\ and Equiv_F :: "'f \ 'f \ bool" (infix "\" 50) and Prec_F :: "'f \ 'f \ bool" (infix "\\" 50) and Prec_L :: "'l \ 'l \ bool" (infix "\L" 50) and active :: 'l begin inductive step :: "'f inference set \ ('f \ 'l) set \ 'f inference set \ ('f \ 'l) set \ bool" (infix "\LGC" 50) where process: "N1 = N \ M \ N2 = N \ M' \ M \ Red_F (N \ M') \ active_subset M' = {} \ (T, N1) \LGC (T, N2)" | schedule_infer: "T2 = T1 \ T' \ N1 = N \ {(C, L)} \ N2 = N \ {(C, active)} \ L \ active \ T' = no_labels.Inf_between (fst ` (active_subset N)) {C} \ (T1, N1) \LGC (T2, N2)" | compute_infer: "T1 = T2 \ {\} \ N2 = N1 \ M \ active_subset M = {} \ \ \ no_labels.Red_I (fst ` (N1 \ M)) \ (T1, N1) \LGC (T2, N2)" | delete_orphans: "T1 = T2 \ T' \ T' \ no_labels.Inf_from (fst ` (active_subset N)) = {} \ (T1, N) \LGC (T2, N)" lemma premise_free_inf_always_from: "\ \ Inf_F \ prems_of \ = [] \ \ \ no_labels.Inf_from N" unfolding no_labels.Inf_from_def by simp lemma one_step_equiv: "(T1, N1) \LGC (T2, N2) \ N1 \L N2" proof (cases "(T1, N1)" "(T2, N2)" rule: step.cases) show "(T1, N1) \LGC (T2, N2) \ (T1, N1) \LGC (T2, N2)" by blast next fix N M M' assume n1_is: "N1 = N \ M" and n2_is: "N2 = N \ M'" and m_red: "M \ Red_F (N \ M')" have "N1 - N2 \ Red_F N2" using n1_is n2_is m_red by auto then show "N1 \L N2" unfolding derive.simps by blast next fix N C L M assume n1_is: "N1 = N \ {(C, L)}" and not_active: "L \ active" and n2_is: "N2 = N \ {(C, active)}" have "(C, active) \ N2" using n2_is by auto moreover have "C \\ C" by (metis equivp_def equiv_equiv_F) moreover have "active \L L" using active_minimal[OF not_active] . ultimately have "{(C, L)} \ Red_F N2" using red_labeled_clauses by blast then have "N1 - N2 \ Red_F N2" using std_Red_F_eq using n1_is n2_is by blast then show "N1 \L N2" unfolding derive.simps by blast next fix M assume n2_is: "N2 = N1 \ M" have "N1 - N2 \ Red_F N2" using n2_is by blast then show "N1 \L N2" unfolding derive.simps by blast next assume n2_is: "N2 = N1" have "N1 - N2 \ Red_F N2" using n2_is by blast then show "N1 \L N2" unfolding derive.simps by blast qed (* lem:lgc-derivations-are-red-derivations *) lemma lgc_to_red: "chain (\LGC) Ns \ chain (\L) (lmap snd Ns)" using one_step_equiv Lazy_List_Chain.chain_mono by (smt chain_lmap prod.collapse) (* lem:fair-lgc-derivations *) lemma lgc_fair: assumes deriv: "chain (\LGC) Ns" and init_state: "active_subset (snd (lhd Ns)) = {}" and final_state: "passive_subset (Liminf_llist (lmap snd Ns)) = {}" and no_prems_init_active: "\\ \ Inf_F. prems_of \ = [] \ \ \ fst (lhd Ns)" and final_schedule: "Liminf_llist (lmap fst Ns) = {}" shows "fair (lmap snd Ns)" unfolding fair_def proof fix \ assume i_in: "\ \ Inf_from (Liminf_llist (lmap snd Ns))" note lhd_is = lhd_conv_lnth[OF chain_not_lnull[OF deriv]] have i_in_inf_fl: "\ \ Inf_FL" using i_in unfolding Inf_from_def by blast have "Liminf_llist (lmap snd Ns) = active_subset (Liminf_llist (lmap snd Ns))" using final_state unfolding passive_subset_def active_subset_def by blast then have i_in2: "\ \ Inf_from (active_subset (Liminf_llist (lmap snd Ns)))" using i_in by simp define m where "m = length (prems_of \)" then have m_def_F: "m = length (prems_of (to_F \))" unfolding to_F_def by simp have i_in_F: "to_F \ \ Inf_F" using i_in Inf_FL_to_Inf_F unfolding Inf_from_def to_F_def by blast have exist_nj: "\j \ {0..nj. enat (Suc nj) < llength Ns \ prems_of \ ! j \ active_subset (snd (lnth Ns nj)) \ (\k. k > nj \ enat k < llength Ns \ prems_of \ ! j \ active_subset (snd (lnth Ns k))))" proof clarify fix j assume j_in: "j \ {0.. ! j" using i_in2 unfolding m_def Inf_from_def active_subset_def by (smt Collect_mem_eq Collect_mono_iff atLeastLessThan_iff nth_mem old.prod.exhaust snd_conv) then have "(C, active) \ Liminf_llist (lmap snd Ns)" using j_in i_in unfolding m_def Inf_from_def by force then obtain nj where nj_is: "enat nj < llength Ns" and c_in2: "(C, active) \ \ (snd ` (lnth Ns ` {k. nj \ k \ enat k < llength Ns}))" unfolding Liminf_llist_def using init_state by fastforce then have c_in3: "\k. k \ nj \ enat k < llength Ns \ (C, active) \ snd (lnth Ns k)" by blast have nj_pos: "nj > 0" using init_state c_in2 nj_is unfolding active_subset_def lhd_is by fastforce obtain nj_min where nj_min_is: "nj_min = (LEAST nj. enat nj < llength Ns \ (C, active) \ \ (snd ` (lnth Ns ` {k. nj \ k \ enat k < llength Ns})))" by blast then have in_allk: "\k. k \ nj_min \ enat k < llength Ns \ (C, active) \ snd (lnth Ns k)" using c_in3 nj_is c_in2 INT_E LeastI_ex by (smt INT_iff INT_simps(10) c_is image_eqI mem_Collect_eq) have njm_smaller_D: "enat nj_min < llength Ns" using nj_min_is by (smt LeastI_ex \\thesis. (\nj. \enat nj < llength Ns; (C, active) \ \ (snd ` (lnth Ns ` {k. nj \ k \ enat k < llength Ns}))\ \ thesis) \ thesis\) have "nj_min > 0" using nj_is c_in2 nj_pos nj_min_is lhd_is by (metis (mono_tags, lifting) active_subset_def emptyE in_allk init_state mem_Collect_eq not_less snd_conv zero_enat_def chain_length_pos[OF deriv]) then obtain njm_prec where nj_prec_is: "Suc njm_prec = nj_min" using gr0_conv_Suc by auto then have njm_prec_njm: "njm_prec < nj_min" by blast then have njm_prec_njm_enat: "enat njm_prec < enat nj_min" by simp have njm_prec_smaller_d: "njm_prec < llength Ns" - using HOL.no_atp(15)[OF njm_smaller_D njm_prec_njm_enat] . + by (rule less_trans[OF njm_prec_njm_enat njm_smaller_D]) have njm_prec_all_suc: "\k>njm_prec. enat k < llength Ns \ (C, active) \ snd (lnth Ns k)" using nj_prec_is in_allk by simp have notin_njm_prec: "(C, active) \ snd (lnth Ns njm_prec)" proof (rule ccontr) assume "\ (C, active) \ snd (lnth Ns njm_prec)" then have absurd_hyp: "(C, active) \ snd (lnth Ns njm_prec)" by simp have prec_smaller: "enat njm_prec < llength Ns" using nj_min_is nj_prec_is by (smt LeastI_ex Suc_leD \\thesis. (\nj. \enat nj < llength Ns; (C, active) \ \ (snd ` (lnth Ns ` {k. nj \ k \ enat k < llength Ns}))\ \ thesis) \ thesis\ enat_ord_simps(1) le_eq_less_or_eq le_less_trans) have "(C, active) \ \ (snd ` (lnth Ns ` {k. njm_prec \ k \ enat k < llength Ns}))" proof - { fix k assume k_in: "njm_prec \ k \ enat k < llength Ns" have "k = njm_prec \ (C, active) \ snd (lnth Ns k)" using absurd_hyp by simp moreover have "njm_prec < k \ (C, active) \ snd (lnth Ns k)" using nj_prec_is in_allk k_in by simp ultimately have "(C, active) \ snd (lnth Ns k)" using k_in by fastforce } then show "(C, active) \ \ (snd ` (lnth Ns ` {k. njm_prec \ k \ enat k < llength Ns}))" by blast qed then have "enat njm_prec < llength Ns \ (C, active) \ \ (snd ` (lnth Ns ` {k. njm_prec \ k \ enat k < llength Ns}))" using prec_smaller by blast then show False using nj_min_is nj_prec_is Orderings.wellorder_class.not_less_Least njm_prec_njm by blast qed then have notin_active_subs_njm_prec: "(C, active) \ active_subset (snd (lnth Ns njm_prec))" unfolding active_subset_def by blast then show "\nj. enat (Suc nj) < llength Ns \ prems_of \ ! j \ active_subset (snd (lnth Ns nj)) \ (\k. k > nj \ enat k < llength Ns \ prems_of \ ! j \ active_subset (snd (lnth Ns k)))" using c_is njm_prec_all_suc njm_prec_smaller_d by (metis (mono_tags, lifting) active_subset_def mem_Collect_eq nj_prec_is njm_smaller_D snd_conv) qed define nj_set where "nj_set = {nj. (\j\{0.. prems_of \ ! j \ active_subset (snd (lnth Ns nj)) \ (\k. k > nj \ enat k < llength Ns \ prems_of \ ! j \ active_subset (snd (lnth Ns k))))}" { assume m_null: "m = 0" then have "enat 0 < llength Ns \ to_F \ \ fst (lhd Ns)" using no_prems_init_active i_in_F m_def_F zero_enat_def chain_length_pos[OF deriv] by auto then have "\n. enat n < llength Ns \ to_F \ \ fst (lnth Ns n)" unfolding lhd_is by blast } moreover { assume m_pos: "m > 0" have nj_not_empty: "nj_set \ {}" proof - have zero_in: "0 \ {0.. ! 0 \ active_subset (snd (lnth Ns n0))" and "\k>n0. enat k < llength Ns \ prems_of \ ! 0 \ active_subset (snd (lnth Ns k))" using exist_nj by fast then have "n0 \ nj_set" unfolding nj_set_def using zero_in by blast then show "nj_set \ {}" by auto qed have nj_finite: "finite nj_set" using all_ex_finite_set[OF exist_nj] by (metis (no_types, lifting) Suc_ile_eq dual_order.strict_implies_order linorder_neqE_nat nj_set_def) have "\n \ nj_set. \nj \ nj_set. nj \ n" using nj_not_empty nj_finite using Max_ge Max_in by blast then obtain n where n_in: "n \ nj_set" and n_bigger: "\nj \ nj_set. nj \ n" by blast then obtain j0 where j0_in: "j0 \ {0.. ! j0 \ active_subset (snd (lnth Ns n))" and j0_allin: "(\k. k > n \ enat k < llength Ns \ prems_of \ ! j0 \ active_subset (snd (lnth Ns k)))" unfolding nj_set_def by blast obtain C0 where C0_is: "prems_of \ ! j0 = (C0, active)" using j0_in i_in2 unfolding m_def Inf_from_def active_subset_def by (smt Collect_mem_eq Collect_mono_iff atLeastLessThan_iff nth_mem old.prod.exhaust snd_conv) then have C0_prems_i: "(C0, active) \ set (prems_of \)" using in_set_conv_nth j0_in m_def by force have C0_in: "(C0, active) \ (snd (lnth Ns (Suc n)))" using C0_is j0_allin suc_n_length by (simp add: active_subset_def) have C0_notin: "(C0, active) \ (snd (lnth Ns n))" using C0_is j0_notin unfolding active_subset_def by simp have step_n: "lnth Ns n \LGC lnth Ns (Suc n)" using deriv chain_lnth_rel n_in unfolding nj_set_def by blast have is_scheduled: "\T2 T1 T' N1 N C L N2. lnth Ns n = (T1, N1) \ lnth Ns (Suc n) = (T2, N2) \ T2 = T1 \ T' \ N1 = N \ {(C, L)} \ N2 = N \ {(C, active)} \ L \ active \ T' = no_labels.Inf_between (fst ` active_subset N) {C}" using step.simps[of "lnth Ns n" "lnth Ns (Suc n)"] step_n C0_in C0_notin unfolding active_subset_def by fastforce then obtain T2 T1 T' N1 N L N2 where nth_d_is: "lnth Ns n = (T1, N1)" and suc_nth_d_is: "lnth Ns (Suc n) = (T2, N2)" and t2_is: "T2 = T1 \ T'" and n1_is: "N1 = N \ {(C0, L)}" "N2 = N \ {(C0, active)}" and l_not_active: "L \ active" and tp_is: "T' = no_labels.Inf_between (fst ` active_subset N) {C0}" using C0_in C0_notin j0_in C0_is using active_subset_def by fastforce have "j \ {0.. prems_of \ ! j \ prems_of \ ! j0 \ prems_of \ ! j \ (active_subset N)" for j proof - fix j assume j_in: "j \ {0.. ! j \ prems_of \ ! j0" obtain nj where nj_len: "enat (Suc nj) < llength Ns" and nj_prems: "prems_of \ ! j \ active_subset (snd (lnth Ns nj))" and nj_greater: "(\k. k > nj \ enat k < llength Ns \ prems_of \ ! j \ active_subset (snd (lnth Ns k)))" using exist_nj j_in by blast then have "nj \ nj_set" unfolding nj_set_def using j_in by blast moreover have "nj \ n" proof (rule ccontr) assume "\ nj \ n" then have "prems_of \ ! j = (C0, active)" using C0_in C0_notin step.simps[of "lnth Ns n" "lnth Ns (Suc n)"] step_n active_subset_def is_scheduled nj_greater nj_prems suc_n_length by auto then show False using j_not_j0 C0_is by simp qed ultimately have "nj < n" using n_bigger by force then have "prems_of \ ! j \ (active_subset (snd (lnth Ns n)))" using nj_greater n_in Suc_ile_eq dual_order.strict_implies_order unfolding nj_set_def by blast then show "prems_of \ ! j \ (active_subset N)" using nth_d_is l_not_active n1_is unfolding active_subset_def by force qed then have prems_i_active: "set (prems_of \) \ active_subset N \ {(C0, active)}" using C0_prems_i C0_is m_def by (metis Un_iff atLeast0LessThan in_set_conv_nth insertCI lessThan_iff subrelI) moreover have "\ (set (prems_of \) \ active_subset N - {(C0, active)})" using C0_prems_i by blast ultimately have "\ \ Inf_between (active_subset N) {(C0, active)}" using i_in_inf_fl prems_i_active unfolding Inf_between_def Inf_from_def by blast then have "to_F \ \ no_labels.Inf_between (fst ` (active_subset N)) {C0}" unfolding to_F_def Inf_between_def Inf_from_def no_labels.Inf_between_def no_labels.Inf_from_def using Inf_FL_to_Inf_F by force then have i_in_t2: "to_F \ \ T2" using tp_is t2_is by simp have "j \ {0.. (\k. k > n \ enat k < llength Ns \ prems_of \ ! j \ active_subset (snd (lnth Ns k)))" for j proof (cases "j = j0") case True assume "j = j0" then show "(\k. k > n \ enat k < llength Ns \ prems_of \ ! j \ active_subset (snd (lnth Ns k)))" using j0_allin by simp next case False assume j_in: "j \ {0.. j0" obtain nj where nj_len: "enat (Suc nj) < llength Ns" and nj_prems: "prems_of \ ! j \ active_subset (snd (lnth Ns nj))" and nj_greater: "(\k. k > nj \ enat k < llength Ns \ prems_of \ ! j \ active_subset (snd (lnth Ns k)))" using exist_nj j_in by blast then have "nj \ nj_set" unfolding nj_set_def using j_in by blast then show "(\k. k > n \ enat k < llength Ns \ prems_of \ ! j \ active_subset (snd (lnth Ns k)))" using nj_greater n_bigger by auto qed then have allj_allk: "(\c\ set (prems_of \). (\k. k > n \ enat k < llength Ns \ c \ active_subset (snd (lnth Ns k))))" using m_def by (metis atLeast0LessThan in_set_conv_nth lessThan_iff) have "\c\ set (prems_of \). snd c = active" using prems_i_active unfolding active_subset_def by auto then have ex_n_i_in: "\n. enat (Suc n) < llength Ns \ to_F \ \ fst (lnth Ns (Suc n)) \ (\c\ set (prems_of \). snd c = active) \ (\c\ set (prems_of \). (\k. k > n \ enat k < llength Ns \ c \ active_subset (snd (lnth Ns k))))" using allj_allk i_in_t2 suc_nth_d_is fstI n_in nj_set_def by auto then have "\n. enat n < llength Ns \ to_F \ \ fst (lnth Ns n) \ (\c\ set (prems_of \). snd c = active) \ (\c\ set (prems_of \). (\k. k \ n \ enat k < llength Ns \ c \ active_subset (snd (lnth Ns k))))" by auto } ultimately obtain n T2 N2 where i_in_suc_n: "to_F \ \ fst (lnth Ns n)" and all_prems_active_after: "m > 0 \ (\c\ set (prems_of \). (\k. k \ n \ enat k < llength Ns \ c \ active_subset (snd (lnth Ns k))))" and suc_n_length: "enat n < llength Ns" and suc_nth_d_is: "lnth Ns n = (T2, N2)" by (metis less_antisym old.prod.exhaust zero_less_Suc) then have i_in_t2: "to_F \ \ T2" by simp have "\p\n. enat (Suc p) < llength Ns \ to_F \ \ (fst (lnth Ns p)) \ to_F \ \ (fst (lnth Ns (Suc p)))" proof (rule ccontr) assume contra: "\ (\p\n. enat (Suc p) < llength Ns \ to_F \ \ (fst (lnth Ns p)) \ to_F \ \ (fst (lnth Ns (Suc p))))" then have i_in_suc: "p0 \ n \ enat (Suc p0) < llength Ns \ to_F \ \ (fst (lnth Ns p0)) \ to_F \ \ (fst (lnth Ns (Suc p0)))" for p0 by blast have "p0 \ n \ enat p0 < llength Ns \ to_F \ \ (fst (lnth Ns p0))" for p0 proof (induction rule: nat_induct_at_least) case base then show ?case using i_in_t2 suc_nth_d_is by simp next case (Suc p0) assume p_bigger_n: "n \ p0" and induct_hyp: "enat p0 < llength Ns \ to_F \ \ fst (lnth Ns p0)" and sucsuc_smaller_d: "enat (Suc p0) < llength Ns" have suc_p_bigger_n: "n \ p0" using p_bigger_n by simp have suc_smaller_d: "enat p0 < llength Ns" using sucsuc_smaller_d Suc_ile_eq dual_order.strict_implies_order by blast then have "to_F \ \ fst (lnth Ns p0)" using induct_hyp by blast then show ?case using i_in_suc[OF suc_p_bigger_n sucsuc_smaller_d] by blast qed then have i_in_all_bigger_n: "\j. j \ n \ enat j < llength Ns \ to_F \ \ (fst (lnth Ns j))" by presburger have "llength (lmap fst Ns) = llength Ns" by force then have "to_F \ \ \ (lnth (lmap fst Ns) ` {j. n \ j \ enat j < llength (lmap fst Ns)})" using i_in_all_bigger_n using Suc_le_D by auto then have "to_F \ \ Liminf_llist (lmap fst Ns)" unfolding Liminf_llist_def using suc_n_length by auto then show False using final_schedule by fast qed then obtain p where p_greater_n: "p \ n" and p_smaller_d: "enat (Suc p) < llength Ns" and i_in_p: "to_F \ \ (fst (lnth Ns p))" and i_notin_suc_p: "to_F \ \ (fst (lnth Ns (Suc p)))" by blast have p_neq_n: "Suc p \ n" using i_notin_suc_p i_in_suc_n by blast have step_p: "lnth Ns p \LGC lnth Ns (Suc p)" using deriv p_smaller_d chain_lnth_rel by blast then have "\T1 T2 \ N2 N1 M. lnth Ns p = (T1, N1) \ lnth Ns (Suc p) = (T2, N2) \ T1 = T2 \ {\} \ N2 = N1 \ M \ active_subset M = {} \ \ \ no_labels.Red_I_\ (fst ` (N1 \ M))" proof - have ci_or_do: "(\T1 T2 \ N2 N1 M. lnth Ns p = (T1, N1) \ lnth Ns (Suc p) = (T2, N2) \ T1 = T2 \ {\} \ N2 = N1 \ M \ active_subset M = {} \ \ \ no_labels.Red_I_\ (fst ` (N1 \ M))) \ (\T1 T2 T' N. lnth Ns p = (T1, N) \ lnth Ns (Suc p) = (T2, N) \ T1 = T2 \ T' \ T' \ no_labels.Inf_from (fst ` active_subset N) = {})" using step.simps[of "lnth Ns p" "lnth Ns (Suc p)"] step_p i_in_p i_notin_suc_p by fastforce then have p_greater_n_strict: "n < Suc p" using suc_nth_d_is p_greater_n i_in_t2 i_notin_suc_p le_eq_less_or_eq by force have "m > 0 \ j \ {0.. prems_of (to_F \) ! j \ fst ` active_subset (snd (lnth Ns p))" for j proof - fix j assume m_pos: "m > 0" and j_in: "j \ {0.. ! j \ (active_subset (snd (lnth Ns p)))" using all_prems_active_after[OF m_pos] p_smaller_d m_def p_greater_n p_neq_n by (meson Suc_ile_eq atLeastLessThan_iff dual_order.strict_implies_order nth_mem p_greater_n_strict) then have "fst (prems_of \ ! j) \ fst ` active_subset (snd (lnth Ns p))" by blast then show "prems_of (to_F \) ! j \ fst ` active_subset (snd (lnth Ns p))" unfolding to_F_def using j_in m_def by simp qed then have prems_i_active_p: "m > 0 \ to_F \ \ no_labels.Inf_from (fst ` active_subset (snd (lnth Ns p)))" using i_in_F unfolding no_labels.Inf_from_def by (smt atLeast0LessThan in_set_conv_nth lessThan_iff m_def_F mem_Collect_eq subsetI) have "m = 0 \ (\T1 T2 \ N2 N1 M. lnth Ns p = (T1, N1) \ lnth Ns (Suc p) = (T2, N2) \ T1 = T2 \ {\} \ N2 = N1 \ M \ active_subset M = {} \ \ \ no_labels.Red_I_\ (fst ` (N1 \ M)))" using ci_or_do premise_free_inf_always_from[of "to_F \" "fst ` active_subset _", OF i_in_F] m_def i_in_p i_notin_suc_p m_def_F by auto then show "(\T1 T2 \ N2 N1 M. lnth Ns p = (T1, N1) \ lnth Ns (Suc p) = (T2, N2) \ T1 = T2 \ {\} \ N2 = N1 \ M \ active_subset M = {} \ \ \ no_labels.Red_I_\ (fst ` (N1 \ M)))" using ci_or_do i_in_p i_notin_suc_p prems_i_active_p unfolding active_subset_def by force qed then obtain T1p T2p N1p N2p Mp where "lnth Ns p = (T1p, N1p)" and suc_p_is: "lnth Ns (Suc p) = (T2p, N2p)" and "T1p = T2p \ {to_F \}" and "T2p \ {to_F \} = {}" and n2p_is: "N2p = N1p \ Mp"and "active_subset Mp = {}" and i_in_red_inf: "to_F \ \ no_labels.Red_I_\ (fst ` (N1p \ Mp))" using i_in_p i_notin_suc_p by fastforce have "to_F \ \ no_labels.Red_I (fst ` (snd (lnth Ns (Suc p))))" using i_in_red_inf suc_p_is n2p_is by fastforce then have "\q \ Q. (\_I_q q (to_F \) \ None \ the (\_I_q q (to_F \)) \ Red_I_q q (\ (\_F_q q ` fst ` snd (lnth Ns (Suc p))))) \ (\_I_q q (to_F \) = None \ \_F_q q (concl_of (to_F \)) \ \ (\_F_q q ` fst ` snd (lnth Ns (Suc p))) \ Red_F_q q (\ (\_F_q q ` fst ` snd (lnth Ns (Suc p)))))" unfolding to_F_def no_labels.Red_I_def no_labels.Red_I_\_q_def by blast then have "\ \ Red_I_\ (snd (lnth Ns (Suc p)))" using i_in_inf_fl unfolding Red_I_\_def Red_I_\_q_def by (simp add: to_F_def) then show "\ \ Sup_llist (lmap Red_I_\ (lmap snd Ns))" unfolding Sup_llist_def using suc_n_length p_smaller_d by auto qed theorem lgc_complete_Liminf: assumes deriv: "chain (\LGC) Ns" and init_state: "active_subset (snd (lhd Ns)) = {}" and final_state: "passive_subset (Liminf_llist (lmap snd Ns)) = {}" and no_prems_init_active: "\\ \ Inf_F. prems_of \ = [] \ \ \ fst (lhd Ns)" and final_schedule: "Liminf_llist (lmap fst Ns) = {}" and b_in: "B \ Bot_F" and bot_entailed: "no_labels.entails_\ (fst ` snd (lhd Ns)) {B}" shows "\BL \ Bot_FL. BL \ Liminf_llist (lmap snd Ns)" proof - have labeled_b_in: "(B, active) \ Bot_FL" using b_in by simp have simp_snd_lmap: "lhd (lmap snd Ns) = snd (lhd Ns)" by (rule llist.map_sel(1)[OF chain_not_lnull[OF deriv]]) have labeled_bot_entailed: "entails_\_L (snd (lhd Ns)) {(B, active)}" using labeled_entailment_lifting bot_entailed by fastforce have "fair (lmap snd Ns)" using lgc_fair[OF deriv init_state final_state no_prems_init_active final_schedule] . then show ?thesis using dynamically_complete_Liminf labeled_b_in lgc_to_red[OF deriv] labeled_bot_entailed simp_snd_lmap std_Red_I_eq by presburger qed (* thm:lgc-completeness *) theorem lgc_complete: assumes deriv: "chain (\LGC) Ns" and init_state: "active_subset (snd (lhd Ns)) = {}" and final_state: "passive_subset (Liminf_llist (lmap snd Ns)) = {}" and no_prems_init_active: "\\ \ Inf_F. prems_of \ = [] \ \ \ fst (lhd Ns)" and final_schedule: "Liminf_llist (lmap fst Ns) = {}" and b_in: "B \ Bot_F" and bot_entailed: "no_labels.entails_\ (fst ` snd (lhd Ns)) {B}" shows "\i. enat i < llength Ns \ (\BL \ Bot_FL. BL \ snd (lnth Ns i))" proof - have "\BL\Bot_FL. BL \ Liminf_llist (lmap snd Ns)" using assms by (rule lgc_complete_Liminf) then show ?thesis unfolding Liminf_llist_def by auto qed end end diff --git a/thys/Virtual_Substitution/OptimizationProofs.thy b/thys/Virtual_Substitution/OptimizationProofs.thy --- a/thys/Virtual_Substitution/OptimizationProofs.thy +++ b/thys/Virtual_Substitution/OptimizationProofs.thy @@ -1,702 +1,702 @@ subsection "Optimization Proofs" theory OptimizationProofs imports Optimizations begin lemma neg_nnf : "\\. (\ eval (nnf (Neg \)) \) = eval (nnf \) \" apply(induction \) apply(simp_all) using aNeg_aEval apply blast using aNeg_aEval by blast theorem eval_nnf : "\\. eval \ \ = eval (nnf \) \" apply(induction \)apply(simp_all) using neg_nnf by blast theorem negation_free_nnf : "negation_free (nnf \)" proof(induction "depth \" arbitrary : \ rule: nat_less_induct ) case 1 then show ?case proof(induction \) case (And \1 \2) then show ?case apply simp by (metis less_Suc_eq_le max.cobounded1 max.cobounded2) next case (Or \1 \2) then show ?case apply simp by (metis less_Suc_eq_le max.cobounded1 max.cobounded2) next case (Neg \) then show ?case proof (induction \) case (And \1 \2) then show ?case apply simp by (metis less_Suc_eq max_less_iff_conj not_less_eq) next case (Or \1 \2) then show ?case apply simp by (metis less_Suc_eq max_less_iff_conj not_less_eq) next case (Neg \) then show ?case by (metis Suc_eq_plus1 add_lessD1 depth.simps(6) lessI nnf.simps(12)) qed auto qed auto qed lemma groupQuantifiers_eval : "eval F L = eval (groupQuantifiers F) L" apply(induction F arbitrary: L rule:groupQuantifiers.induct) unfolding doubleExist unwrapExist unwrapExist' unwrapExist'' doubleForall unwrapForall unwrapForall' unwrapForall'' apply (auto) using doubleExist doubleExist unwrapExist unwrapExist' unwrapExist'' doubleForall unwrapForall unwrapForall' unwrapForall'' apply auto by metis+ theorem simp_atom_eval : "aEval a xs = eval (simp_atom a) xs" proof(cases a) case (Less p) then show ?thesis by(cases "get_if_const p")(simp_all add:get_if_const_insertion) next case (Eq p) then show ?thesis by(cases "get_if_const p")(simp_all add:get_if_const_insertion) next case (Leq p) then show ?thesis by(cases "get_if_const p")(simp_all add:get_if_const_insertion) next case (Neq p) then show ?thesis by(cases "get_if_const p")(simp_all add:get_if_const_insertion) qed lemma simpfm_eval : "\L. eval \ L = eval (simpfm \) L" apply(induction \) apply(simp_all add: simp_atom_eval eval_and eval_or) using eval_neg by blast lemma exQ_clearQuantifiers: assumes ExQ : "\xs. eval (clearQuantifiers \) xs = eval \ xs" shows "eval (clearQuantifiers (ExQ \)) xs = eval (ExQ \) xs" proof- define \' where "\' = clearQuantifiers \" have h : "freeIn 0 \' \ (eval (lowerFm 0 1 \') xs = eval (ExQ \') xs)" using eval_lowerFm by simp have "eval (clearQuantifiers (ExQ \)) xs = eval (if freeIn 0 \' then lowerFm 0 1 \' else ExQ \') xs" using \'_def by simp also have "... = eval (ExQ \) xs" apply(cases "freeIn 0 \'") using h ExQ \'_def by(simp_all) finally show ?thesis by simp qed lemma allQ_clearQuantifiers : assumes AllQ : "\xs. eval (clearQuantifiers \) xs = eval \ xs" shows "eval (clearQuantifiers (AllQ \)) xs = eval (AllQ \) xs" proof- define \' where "\' = clearQuantifiers \" have "freeIn 0 \' \ (eval (ExQ \') xs) = eval (AllQ \') xs" by (simp add: var_not_in_eval2) then have h : "freeIn 0 \' \ (eval (lowerFm 0 1 \') xs = eval (AllQ \') xs)" using eval_lowerFm by simp have "eval (clearQuantifiers (AllQ \)) xs = eval (if freeIn 0 \' then lowerFm 0 1 \' else AllQ \') xs" using \'_def by simp also have "... = eval (AllQ \) xs" apply(cases "freeIn 0 \'") using h AllQ \'_def by(simp_all) finally show ?thesis by simp qed lemma clearQuantifiers_eval : "eval (clearQuantifiers \) xs = eval \ xs" proof(induction \ arbitrary : xs) case (Atom x) then show ?case using simp_atom_eval by simp next case (And \1 \2) then show ?case using eval_and by simp next case (Or \1 \2) then show ?case using eval_or by simp next case (Neg \) then show ?case using eval_neg by auto next case (ExQ \) then show ?case using exQ_clearQuantifiers by simp next case (AllQ \) then show ?case using allQ_clearQuantifiers by simp next case (ExN x1 \) then show ?case proof(induction x1 arbitrary:\) case 0 then show ?case by auto next case (Suc x1) show ?case using Suc(1)[of "ExQ \", OF exQ_clearQuantifiers[OF Suc(2)]] apply simp using Suc_eq_plus1 \eval (clearQuantifiers (ExN x1 (ExQ \))) xs = eval (ExN x1 (ExQ \)) xs\ eval.simps(10) unwrapExist' by presburger qed next case (AllN x1 \) then show ?case proof(induction x1 arbitrary:\) case 0 then show ?case by auto next case (Suc x1) show ?case using Suc(1)[of "AllQ \", OF allQ_clearQuantifiers[OF Suc(2)]] apply simp using unwrapForall' by force qed qed auto lemma push_forall_eval_AllQ : "\xs. eval (AllQ \) xs = eval (push_forall (AllQ \)) xs" proof(induction \) case TrueF then show ?case by simp next case FalseF then show ?case by simp next case (Atom x) then show ?case using aEval_lowerAtom eval.simps(1) eval.simps(8) push_forall.simps(11) by presburger next case (And \1 \2) {fix xs have "eval (AllQ (And \1 \2)) xs = (\x. eval \1 (x#xs) \ eval \2 (x#xs))" by simp also have "... = ((\x. eval \1 (x#xs)) \ (\x. eval \2 (x#xs)))" by blast also have "... = eval (push_forall (AllQ (And \1 \2))) xs" using And eval_and by(simp) finally have "eval (AllQ (And \1 \2)) xs = eval (push_forall (AllQ (And \1 \2))) xs" by simp } then show ?case by simp next case (Or \1 \2) then show ?case proof(cases "freeIn 0 \1") case True then have h : "freeIn 0 \1" by simp then show ?thesis proof(cases "freeIn 0 \2") case True {fix xs have "\x. eval \1 (x # xs) = eval (lowerFm 0 1 \1) xs" using eval_lowerFm h eval.simps(7) by blast then have h1 : "\x. eval \1 (x # xs) = eval (lowerFm 0 1 \1) xs" using h var_not_in_eval2 by blast have "\x. eval \2 (x # xs) = eval (lowerFm 0 1 \2) xs" using eval_lowerFm True eval.simps(7) by blast then have h2 : "\x. eval \2 (x # xs) = eval (lowerFm 0 1 \2) xs" using True var_not_in_eval2 by blast have "eval (AllQ (Or \1 \2)) xs = eval (push_forall (AllQ (Or \1 \2))) xs" by(simp add:h h1 h2 True eval_or) } then show ?thesis by simp next case False {fix xs have "\x. eval \1 (x # xs) = eval (lowerFm 0 1 \1) xs" using eval_lowerFm h eval.simps(7) by blast then have "\x. eval \1 (x # xs) = eval (lowerFm 0 1 \1) xs" using True var_not_in_eval2 by blast then have "eval (AllQ (Or \1 \2)) xs = eval (push_forall (AllQ (Or \1 \2))) xs" by(simp add:h False eval_or) } then show ?thesis by simp qed next case False then have h : "\freeIn 0 \1" by simp then show ?thesis proof(cases "freeIn 0 \2") case True {fix xs have "\x. eval \2 (x # xs) = eval (lowerFm 0 1 \2) xs" using eval_lowerFm True eval.simps(7) by blast then have "\x. eval \2 (x # xs) = eval (lowerFm 0 1 \2) xs" using True var_not_in_eval2 by blast then have "eval (AllQ (Or \1 \2)) xs = eval (push_forall (AllQ (Or \1 \2))) xs" by(simp add:h True eval_or) } then show ?thesis by simp next case False then show ?thesis by(simp add:h False eval_or) qed qed next case (Neg \) {fix xs have "freeIn 0 (Neg \) \ (eval (ExQ (Neg \)) xs) = eval (AllQ (Neg \)) xs" using var_not_in_eval2 eval.simps(7) eval.simps(8) by blast then have h : "freeIn 0 (Neg \) \ (eval (lowerFm 0 1 (Neg \)) xs = eval (AllQ (Neg \)) xs)" using eval_lowerFm by blast have "eval (push_forall (AllQ (Neg \))) xs = eval (if freeIn 0 (Neg \) then lowerFm 0 1 (Neg \) else AllQ (Neg \)) xs" by simp also have "... = eval (AllQ (Neg \)) xs" apply(cases "freeIn 0 (Neg \)") using h by(simp_all) finally have "eval (push_forall (AllQ (Neg \))) xs = eval (AllQ (Neg \)) xs" by simp } then show ?case by simp next case (ExQ \) {fix xs have "freeIn 0 (ExQ \) \ (eval (ExQ (ExQ \)) xs) = eval (AllQ (ExQ \)) xs" using var_not_in_eval2 eval.simps(7) eval.simps(8) by blast then have h : "freeIn 0 (ExQ \) \ (eval (lowerFm 0 1 (ExQ \)) xs = eval (AllQ (ExQ \)) xs)" using eval_lowerFm by blast have "eval (push_forall (AllQ (ExQ \))) xs = eval (if freeIn 0 (ExQ \) then lowerFm 0 1 (ExQ \) else AllQ (ExQ \)) xs" by simp also have "... = eval (AllQ (ExQ \)) xs" apply(cases "freeIn 0 (ExQ \)") using h by(simp_all) finally have "eval (push_forall (AllQ (ExQ \))) xs = eval (AllQ (ExQ \)) xs" by simp } then show ?case by simp next case (AllQ \) {fix xs have "freeIn 0 (AllQ \) \ (eval (ExQ (AllQ \)) xs) = eval (AllQ (AllQ \)) xs" using var_not_in_eval2 eval.simps(7) eval.simps(8) by blast then have h : "freeIn 0 (AllQ \) \ (eval (lowerFm 0 1 (AllQ \)) xs = eval (AllQ (AllQ \)) xs)" using eval_lowerFm by blast have "eval (push_forall (AllQ (AllQ \))) xs = eval (if freeIn 0 (AllQ \) then lowerFm 0 1 (AllQ \) else AllQ (AllQ \)) xs" by simp also have "... = eval (AllQ (AllQ \)) xs" apply(cases "freeIn 0 (AllQ \)") using h AllQ by(simp_all) finally have "eval (push_forall (AllQ (AllQ \))) xs = eval (AllQ (AllQ \)) xs" by simp } then show ?case by simp next case (ExN x1 \) then show ?case using eval.simps(7) eval.simps(8) eval_lowerFm push_forall.simps(17) var_not_in_eval2 by presburger next case (AllN x1 \) then show ?case using eval.simps(7) eval.simps(8) eval_lowerFm push_forall.simps(18) var_not_in_eval2 by presburger qed lemma push_forall_eval : "\xs. eval \ xs = eval (push_forall \) xs" proof(induction \) case (Atom x) then show ?case using simp_atom_eval by simp next case (And \1 \2) then show ?case using eval_and by auto next case (Or \1 \2) then show ?case using eval_or by auto next case (Neg \) then show ?case using eval_neg by auto next case (AllQ \) then show ?case using push_forall_eval_AllQ by blast next case (ExN x1 \) then show ?case using eval.simps(10) push_forall.simps(7) by presburger qed auto lemma map_fm_binders_negation_free : assumes "negation_free \" shows "negation_free (map_fm_binders f \ n)" using assms apply(induction \ arbitrary : n) by auto lemma negation_free_and : assumes "negation_free \" assumes "negation_free \" shows "negation_free (and \ \)" using assms unfolding and_def by simp lemma negation_free_or : assumes "negation_free \" assumes "negation_free \" shows "negation_free (or \ \)" using assms unfolding or_def by simp lemma push_forall_negation_free_all : assumes "negation_free \" shows "negation_free (push_forall (AllQ \))" using assms proof(induction \) case (And \1 \2) show ?case apply auto apply(rule negation_free_and) using And by auto next case (Or \1 \2) show ?case apply auto apply(rule negation_free_or) using Or map_fm_binders_negation_free negation_free_or by auto next case (ExQ \) then show ?case using map_fm_binders_negation_free by auto next case (AllQ \) then show ?case using map_fm_binders_negation_free by auto next case (ExN x1 \) then show ?case using map_fm_binders_negation_free by auto next case (AllN x1 \) then show ?case using map_fm_binders_negation_free by auto qed auto lemma push_forall_negation_free : assumes "negation_free \" shows "negation_free(push_forall \)" using assms proof(induction \) case (Atom A) then show ?case apply(cases A) by auto next case (And \1 \2) then show ?case by (auto simp add: and_def) next case (Or \1 \2) then show ?case by (auto simp add: or_def) next case (AllQ \) then show ?case using push_forall_negation_free_all by auto qed auto lemma to_list_insertion: "insertion f p = sum_list [insertion f term * (f v) ^ i. (term,i)\(to_list v p)]" proof- have h1 : "insertion f p = insertion f (\i\MPoly_Type.degree p v. isolate_variable_sparse p v i * Var v ^ i)" using sum_over_zero by auto have h2 : "insertion f (Var v) = f v" by (auto simp: monomials_Var coeff_Var insertion_code) define d where "d = MPoly_Type.degree p v" define g where "g = (\x. insertion f (isolate_variable_sparse p v x) * f v ^ x)" have h3 : "insertion f (isolate_variable_sparse p v d) * f v ^ d = g d" using g_def by auto show ?thesis unfolding h1 insertion_sum' insertion_mult insertion_pow h2 apply auto unfolding d_def[symmetric] g_def[symmetric] h3 proof(induction d) case 0 then show ?case by auto next case (Suc d) show ?case apply (auto simp add: Suc ) unfolding g_def by auto qed qed lemma to_list_p: "p = sum_list [term * (Var v) ^ i. (term,i)\(to_list v p)]" proof- define d where "d = MPoly_Type.degree p v" have "(\i\MPoly_Type.degree p v. isolate_variable_sparse p v i * Var v ^ i) = (\(term, i)\to_list v p. term * Var v ^ i)" unfolding to_list.simps d_def[symmetric] apply(induction d) by auto then show ?thesis using sum_over_zero[of p v] by auto qed fun chophelper :: "(real mpoly * nat) list \ (real mpoly * nat) list \ (real mpoly * nat) list * (real mpoly * nat) list" where "chophelper [] L = (L,[])"| "chophelper ((p,i)#L) R = (if p=0 then chophelper L (R @ [(p,i)]) else (R,(p,i)#L))" lemma preserve : assumes "(a,b)=chophelper L L'" shows "a@b=L'@L" using assms proof(induction L arbitrary : a b L') case Nil then show ?case using assms by auto next case (Cons A L) then show ?case proof(cases A) case (Pair p i) show ?thesis using Cons unfolding Pair apply(cases "p=0") by auto qed qed lemma compare : assumes "(a,b)=chophelper L L'" shows "chop L = b" using assms proof(induction L arbitrary : a b L') case Nil then show ?case by auto next case (Cons A L) then show ?case proof(cases A) case (Pair p i) show ?thesis using Cons unfolding Pair apply(cases "p=0") by auto qed qed lemma allzero: assumes "\(p,i)\set(L'). p=0" assumes "(a,b)=chophelper L L'" shows "\(p,i)\set(a). p=0" using assms proof(induction L arbitrary : a b L') case Nil then show ?case by auto next case (Cons t L) then show ?case proof(cases t) case (Pair p i) show ?thesis proof(cases "p=0") case True have h1: "\x\set (L' @ [(0, i)]). case x of (p, i) \ p = 0" using Cons(2) by auto show ?thesis using Cons(1)[OF h1] Cons(3) True unfolding Pair by auto next case False then show ?thesis using Cons unfolding Pair by auto qed qed qed lemma separate: assumes "(a,b)=chophelper (to_list v p) []" shows "insertion f p = sum_list [insertion f term * (f v) ^ i. (term,i)\a] + sum_list [insertion f term * (f v) ^ i. (term,i)\b]" using to_list_insertion[of f p v] preserve[OF assms, symmetric] unfolding List.append.left_neutral by (simp del: to_list.simps) lemma chopped : assumes "(a,b)=chophelper (to_list v p) []" shows "insertion f p = sum_list [insertion f term * (f v) ^ i. (term,i)\b]" proof- have h1 : "\(p, i)\set []. p = 0" by auto have "(\(term, i)\a. insertion f term * f v ^ i) = 0" using allzero[OF h1 assms] proof(induction a) case Nil then show ?case by auto next case (Cons a1 a2) then show ?case apply(cases a1) by simp qed then show ?thesis using separate[OF assms, of f] by auto qed lemma insertion_chop : shows "insertion f p = sum_list [insertion f term * (f v) ^ i. (term,i)\(chop (to_list v p))]" proof(cases "chophelper (to_list v p) []") case (Pair a b) show ?thesis using chopped[OF Pair[symmetric], of f] unfolding compare[OF Pair[symmetric], symmetric] . qed lemma sorted : "sorted_wrt (\(_,i).\(_,i'). ix. (isolate_variable_sparse p v x, x)) [0..x. (isolate_variable_sparse p v x, x)) [0..(chop (to_list v p))] * (f v)^i" proof- have h : "sorted_wrt (\(_, i) (_, y). i < y) (chop (to_list v p))" proof- define L where "L = to_list v p" show ?thesis using sublist[of "to_list v p"] sorted[of v p] unfolding L_def[symmetric] by (metis sorted_wrt_append sublist_def) qed then have "\(term,d)\set(chop (to_list v p)). d\i" unfolding assms[symmetric] by fastforce then have simp : "\(term,d)\set(chop(to_list v p)). f v ^ (d - i) * f v ^ i = f v ^ d" - unfolding HOL.no_atp(118) by(auto simp del: to_list.simps) + unfolding semiring_normalization_rules(26) by(auto simp del: to_list.simps) have "insertion f p = sum_list [insertion f term * (f v) ^ i. (term,i)\(chop (to_list v p))]" using insertion_chop[of f p v] . also have "...= (\(term, d)\chop (to_list v p). insertion f term * f v ^ (d-i) * f v ^ i)" using simp by (smt Pair_inject case_prodE map_eq_conv mult.assoc split_cong) also have "... = (\(term, d)\chop (to_list v p). insertion f term * f v ^ (d - i)) * f v ^ i" proof- define d where "d = chop(to_list v p)" define a where "a = f v ^ i" define b where "b = (\(term, d). insertion f term * f v ^ (d - i))" have h : "(\(term, d)\d. insertion f term * f v ^ (d - i) * a) = (\(term, d)\d. b (term,d) * a)" using b_def by auto show ?thesis unfolding d_def[symmetric] a_def[symmetric] b_def[symmetric] h apply(induction d) apply simp apply auto by (simp add: ring_class.ring_distribs(2)) qed finally show ?thesis by auto qed lemma insert_Var_Zero : "insertion f (Var v) = f v" unfolding insertion_code monomials_Var apply auto unfolding coeff_Var by simp lemma decreasePower_insertion : assumes "decreasePower v p = (p',i)" shows "insertion f p = insertion f p'* (f v)^i" proof(cases "chop (to_list v p)") case Nil then show ?thesis using assms by auto next case (Cons a list) then show ?thesis proof(cases a) case (Pair coef i') have i'_def : "i'=i" using Cons assms Pair by auto have chop: "chop (to_list v p) = (coef, i) # list" using Cons assms unfolding i'_def Pair by auto have p'_def : "p' = (\(term, x)\chop (to_list v p). term * Var v ^ (x - i))" using assms Cons Pair by auto have p'_insertion : "insertion f p' = (\(term, x)\chop (to_list v p). insertion f term * f v ^ (x - i))" proof- define d where "d = chop (to_list v p)" have "insertion f p' = insertion f (\(term, x)\chop (to_list v p). term * Var v ^ (x - i))" using p'_def by auto also have "... = (\(term, x)\chop (to_list v p). insertion f (term * Var v ^ (x - i)))" unfolding d_def[symmetric] apply(induction d) apply simp apply(simp add:insertion_add) by auto also have "... = (\(term, x)\chop (to_list v p). insertion f term * f v ^ (x - i))" unfolding insertion_mult insertion_pow insert_Var_Zero by auto finally show ?thesis by auto qed have h : "(coef, i') # list = chop (to_list v p)" using Cons unfolding Pair by auto show ?thesis unfolding p'_insertion using move_exp[OF h, of f] unfolding i'_def . qed qed lemma unpower_eval: "eval (unpower v \) L = eval \ L" proof(induction \ arbitrary: v L) case TrueF then show ?case by auto next case FalseF then show ?case by auto next case (Atom At) then show ?case proof(cases At) case (Less p) obtain q i where h: "decreasePower v p = (q, i)" using prod.exhaust_sel by blast have p : "\f. insertion f p = insertion f q* (f v)^i" using decreasePower_insertion[OF h] by auto show ?thesis proof(cases "i=0") case True then show ?thesis unfolding Less unpower.simps h by auto next case False obtain x where x_def : "Suc x = i" using False using not0_implies_Suc by auto have h1 : "i mod 2 = 0 \ (insertion (nth_default 0 L) q < 0 \ insertion (nth_default 0 L) (Var v) \ 0) = (insertion (nth_default 0 L) q * nth_default 0 L v ^ i < 0)" proof - assume "i mod 2 = 0" then have "\r. \ (r::real) ^ i < 0" by presburger then show ?thesis by (metis \\thesis. (\x. Suc x = i \ thesis) \ thesis\ insert_Var_Zero linorder_neqE_linordered_idom mult_less_0_iff power_0_Suc power_eq_0_iff) qed show ?thesis unfolding Less unpower.simps h x_def[symmetric] apply simp unfolding x_def p apply(cases "i mod 2 = 0") using h1 apply simp_all by (smt insert_Var_Zero insertion_neg mod_Suc mod_eq_0D mult_less_0_iff nat.inject odd_power_less_zero power_0 power_Suc0_right power_eq_0_iff x_def zero_less_Suc zero_less_power) qed next case (Eq p) obtain q i where h: "decreasePower v p = (q, i)" using prod.exhaust_sel by blast have p : "\f. insertion f p = insertion f q* (f v)^i" using decreasePower_insertion[OF h] by auto show ?thesis unfolding Eq unpower.simps h apply simp apply(cases i) apply simp apply simp unfolding p apply simp by (metis insert_Var_Zero) next case (Leq p) obtain q i where h: "decreasePower v p = (q, i)" using prod.exhaust_sel by blast have p : "\f. insertion f p = insertion f q* (f v)^i" using decreasePower_insertion[OF h] by auto show ?thesis proof(cases "i=0") case True then show ?thesis unfolding Leq unpower.simps h by auto next case False obtain x where x_def : "Suc x = i" using False using not0_implies_Suc by auto define a where "a = insertion (nth_default 0 L) q" define x' where "x' = nth_default 0 L v" show ?thesis unfolding Leq unpower.simps h x_def[symmetric] apply simp unfolding x_def p apply(cases "i mod 2 = 0") unfolding insert_Var_Zero insertion_mult insertion_pow insertion_neg apply simp_all unfolding a_def[symmetric] x'_def[symmetric] proof- assume "i mod 2 = 0" then have "x' ^ i \0" by (simp add: \i mod 2 = 0\ even_iff_mod_2_eq_zero zero_le_even_power) then show "(a \ 0 \ x' = 0) = (a * x' ^ i \ 0)" using Rings.ordered_semiring_0_class.mult_nonpos_nonneg[of a "x'^i"] apply auto unfolding Rings.linordered_ring_strict_class.mult_le_0_iff apply auto by (simp add: False power_0_left) next assume h: "i mod 2 = Suc 0" show "(a = 0 \ a < 0 \ 0 \ x' \ 0 < a \ x' \ 0) = (a * x' ^ i \ 0)" using h by (smt even_iff_mod_2_eq_zero mult_less_cancel_right mult_neg_neg mult_nonneg_nonpos mult_pos_pos not_mod2_eq_Suc_0_eq_0 power_0_Suc x_def zero_le_power_eq zero_less_mult_pos2 zero_less_power) qed qed next case (Neq p) obtain q i where h: "decreasePower v p = (q, i)" using prod.exhaust_sel by blast have p : "\f. insertion f p = insertion f q* (f v)^i" using decreasePower_insertion[OF h] by auto show ?thesis unfolding Neq unpower.simps h apply simp apply(cases i) apply simp apply simp unfolding p apply simp by (metis insert_Var_Zero) qed qed auto lemma to_list_filter: "p = sum_list [term * (Var v) ^ i. (term,i)\((filter (\(x,_). x\0) (to_list v p)))]" proof- define L where "L = to_list v p" have "(\(term, i)\to_list v p. term * Var v ^ i) = (\(term, i)\filter (\(x, _). x \ 0) (to_list v p). term * Var v ^ i)" unfolding L_def[symmetric] apply(induction L) by auto then show ?thesis using to_list_p[of p v] by auto qed end \ No newline at end of file