diff --git a/thys/Aristotles_Assertoric_Syllogistic/AristotlesAssertoric.thy b/thys/Aristotles_Assertoric_Syllogistic/AristotlesAssertoric.thy --- a/thys/Aristotles_Assertoric_Syllogistic/AristotlesAssertoric.thy +++ b/thys/Aristotles_Assertoric_Syllogistic/AristotlesAssertoric.thy @@ -1,247 +1,247 @@ (* Title: Aristotle's Assertoric Syllogistic Author: Angeliki Koutsoukou-Argyraki, University of Cambridge. October 2019 We formalise with Isabelle/HOL some basic elements of Aristotle's assertoric syllogistic following the article from the Stanford Encyclopedia of Philosophy by Robin Smith: https://plato.stanford.edu/entries/aristotle-logic/. To this end, we use a set theoretic formulation (covering both individual and general predication). In particular, we formalise the deductions in the Figures and after that we present Aristotle's metatheoretical observation that all deductions in the Figures can in fact be reduced to either Barbara or Celarent. As the formal proofs prove to be straightforward, the interest of this entry lies in illustrating the functionality of Isabelle and high efficiency of Sledgehammer for simple exercises in philosophy.*) section\Aristotle's Assertoric Syllogistic\ theory AristotlesAssertoric imports Main begin subsection\Aristotelean Categorical Sentences\ text\ Aristotle's universal, particular and indefinite predications (affirmations and denials) are expressed here using a set theoretic formulation. Aristotle handles in the same way individual and general predications i.e. he gives the same logical analysis to "Socrates is an animal" and "humans are animals". Here we define the general predication i.e. predications are defined as relations between sets. This has the benefit that individual predication can also be expressed as set membership (e.g. see the lemma SocratesMortal). \ definition universal_affirmation :: "'a set \'a set \ bool" (infixr "Q" 80) where "A Q B \ \ b \ B . b \ A " definition universal_denial :: "'a set \'a set \ bool" (infixr "E" 80) where "A E B \ \ b \ B. ( b \ A) " definition particular_affirmation :: " 'a set \'a set \ bool" (infixr "I" 80) where "A I B \ \ b \ B. ( b \ A) " definition particular_denial :: "'a set \'a set \ bool" (infixr "Z" 80) where "A Z B \ \ b \ B. ( b \ A) " text\ The above four definitions are known as the "square of opposition".\ definition indefinite_affirmation :: " 'a set \'a set \ bool" (infixr "QI" 80) where "A QI B \(( \ b \ B. (b \ A)) \ (\ b \ B. (b \ A))) " definition indefinite_denial :: "'a set \'a set \ bool" (infixr "EZ" 80) where "A EZ B \ (( \ b \ B. (b \ A)) \ (\ b \ B. (b \ A))) " lemma aristo_conversion1 : assumes "A E B" shows "B E A" using assms universal_denial_def by blast lemma aristo_conversion2 : assumes "A I B" shows "B I A" using assms unfolding particular_affirmation_def by blast lemma aristo_conversion3 : assumes "A Q B" and "B \{} " shows "B I A" using assms unfolding universal_affirmation_def particular_affirmation_def by blast text\Remark: Aristotle in general supposes that sets have to be nonempty. Indeed, we observe that in many instances it is necessary to assume that the sets are nonempty, otherwise Isabelle's automation finds counterexamples.\ subsection\The Deductions in the Figures ("Moods")\ text\The medieval mnemonic names are used.\ subsubsection\First Figure\ lemma Barbara: assumes "A Q B " and "B Q C" shows "A Q C" by (meson assms universal_affirmation_def) lemma Celarent: assumes "A E B " and "B Q C" shows "A E C" by (meson assms universal_affirmation_def universal_denial_def) lemma Darii: assumes "A Q B" and "B I C" shows "A I C" by (meson assms particular_affirmation_def universal_affirmation_def) lemma Ferio: assumes "A E B" and "B I C" shows "A Z C" by (meson assms particular_affirmation_def particular_denial_def universal_denial_def) subsubsection\Second Figure\ lemma Cesare: assumes "A E B " and "A Q C" shows "B E C" using Celarent aristo_conversion1 assms by blast lemma Camestres: assumes "A Q B " and "A E C" shows "B E C " using Cesare aristo_conversion1 assms by blast lemma Festino: assumes "A E B " and "A I C" shows "B Z C " using Ferio aristo_conversion1 assms by blast lemma Baroco: assumes "A Q B " and "A Z C" shows "B Z C " by (meson assms particular_denial_def universal_affirmation_def) subsubsection\Third Figure\ lemma Darapti: assumes "A Q C " and "B Q C" and "C \{}" shows "A I B " using Darii assms unfolding universal_affirmation_def particular_affirmation_def by blast lemma Felapton: assumes "A E C" and "B Q C" and "C \{}" shows "A Z B" using Festino aristo_conversion1 aristo_conversion3 assms by blast lemma Disamis: assumes "A I C" and "B Q C" shows "A I B" using Darii aristo_conversion2 assms by blast lemma Datisi: assumes "A Q C" and "B I C" shows "A I B" using Disamis aristo_conversion2 assms by blast lemma Bocardo: assumes "A Z C" and "B Q C" shows "A Z B" by (meson assms particular_denial_def universal_affirmation_def) lemma Ferison: assumes "A E C " and "B I C" shows "A Z B " using Ferio aristo_conversion2 assms by blast subsubsection\Examples\ text\Example of a deduction with general predication.\ lemma GreekMortal : assumes "Mortal Q Human" and "Human Q Greek " shows " Mortal Q Greek " using assms Barbara by auto text\Example of a deduction with individual predication.\ lemma SocratesMortal: assumes "Socrates \ Human " and "Mortal Q Human" shows "Socrates \ Mortal " using assms by (simp add: universal_affirmation_def) subsection\Metatheoretical comments\ text\The following are presented to demonstrate one of Aristotle's metatheoretical explorations. Namely, Aristotle's metatheorem that: "All deductions in all three Figures can eventually be reduced to either Barbara or Celarent" is demonstrated by the proofs below and by considering the proofs from the previous subsection. \ lemma Darii_reducedto_Camestres: assumes "A Q B " and "B I C" and "A E C " (*assms, concl. of Darii and A E C *) shows "A I C" proof- have "B E C" using Camestres \ A Q B \ \A E C\ by blast show ?thesis using \ B I C \ \B E C\ by (simp add: particular_affirmation_def universal_denial_def) qed text\It is already evident from the proofs in the previous subsection that: Camestres can be reduced to Cesare. Cesare can be reduced to Celarent. Festino can be reduced to Ferio.\ lemma Ferio_reducedto_Cesare: assumes "A E B " and "B I C" and "A Q C " (*assms, concl. of Ferio and A Q C *) shows "A Z C" proof- have "B E C" using Cesare \A E B \ \A Q C\ by blast show ?thesis using \B I C \ \B E C\ by (simp add: particular_affirmation_def universal_denial_def) qed lemma Baroco_reducedto_Barbara : assumes "A Q B " and " A Z C " and " B Q C " shows "B Z C" (*assms , concl. of Baroco and B Q C *) proof- have "A Q C" using \A Q B \ \ B Q C \ Barbara by blast show ?thesis using \A Q C\ \ A Z C \ by (simp add: particular_denial_def universal_affirmation_def) qed lemma Bocardo_reducedto_Barbara : assumes " A Z C" and "B Q C" and "A Q B" shows "A Z B" (*assms, concl of Bocardo and A Q B *) proof- have "A Q C" using \B Q C\ \ A Q B\ using Barbara by blast show ?thesis using \A Q C\ \ A Z C\ by (simp add: particular_denial_def universal_affirmation_def) qed text\Finally, it is already evident from the proofs in the previous subsection that : Darapti can be reduced to Darii. Felapton can be reduced to Festino. Disamis can be reduced to Darii. Datisi can be reduced to Disamis. Ferison can be reduced to Ferio. \ text\In conclusion, the aforementioned deductions have thus been shown to be reduced to either Barbara or Celarent as follows: Baroco $\Rightarrow$ Barbara Bocardo $\Rightarrow$ Barbara Felapton $\Rightarrow$ Festino $\Rightarrow$ Ferio $\Rightarrow$ Cesare $\Rightarrow$ Celarent Datisi $\Rightarrow$ Disamis $\Rightarrow$ Darii $\Rightarrow$ Camestres $\Rightarrow$ Cesare Darapti $\Rightarrow$ Darii Ferison $\Rightarrow$ Ferio \ subsection\Acknowledgements\ text\A.K.-A. was supported by the ERC Advanced Grant ALEXANDRIA (Project 742178) funded by the European Research Council and led by Professor Lawrence Paulson at the University of Cambridge, UK. Thanks to Wenda Li.\ subsection\Bibliography\ -text\Smith, Robin, "Aristotle’s Logic", +text\Smith, Robin, "Aristotle's Logic", The Stanford Encyclopedia of Philosophy (Summer 2019 Edition), Edward N. Zalta (ed.), URL = @{url "https://plato.stanford.edu/archives/sum2019/entries/aristotle-logic/"} \ end diff --git a/thys/Attack_Trees/GDPRhealthcare.thy b/thys/Attack_Trees/GDPRhealthcare.thy --- a/thys/Attack_Trees/GDPRhealthcare.thy +++ b/thys/Attack_Trees/GDPRhealthcare.thy @@ -1,387 +1,387 @@ section \Application example from IoT healthcare\ text \The example of an IoT healthcare systems is taken from the context of the CHIST-ERA project SUCCESS \cite{suc:16}. In this system architecture, data is collected by sensors in the home or via a smart phone helping to monitor bio markers of the patient. The data collection is in a cloud based server to enable hospitals (or scientific institutions) to access the data which is controlled via the smart phone. The identities Patient and Doctor represent patients and their doctors; double quotes ''s'' indicate strings in Isabelle/HOL. The global policy is `only the patient and the doctor can access the data in the cloud'.\ theory GDPRhealthcare imports Infrastructure begin text \Local policies are represented as a function over an @{text \igraph G\} that additionally assigns each location of a scenario to its local policy given as a pair of requirements to an actor (first element of the pair) in order to grant him actions in the location (second element of the pair). The predicate @{text \@G\} checks whether an actor is at a given location in the @{text \igraph G\}.\ locale scenarioGDPR = fixes gdpr_actors :: "identity set" defines gdpr_actors_def: "gdpr_actors \ {''Patient'', ''Doctor''}" fixes gdpr_locations :: "location set" defines gdpr_locations_def: "gdpr_locations \ {Location 0, Location 1, Location 2, Location 3}" fixes sphone :: "location" defines sphone_def: "sphone \ Location 0" fixes home :: "location" defines home_def: "home \ Location 1" fixes hospital :: "location" defines hospital_def: "hospital \ Location 2" fixes cloud :: "location" defines cloud_def: "cloud \ Location 3" fixes global_policy :: "[infrastructure, identity] \ bool" defines global_policy_def: "global_policy I a \ a \ ''Doctor'' \ \(enables I hospital (Actor a) eval)" fixes global_policy' :: "[infrastructure, identity] \ bool" defines global_policy'_def: "global_policy' I a \ a \ gdpr_actors \ \(enables I cloud (Actor a) get)" fixes ex_creds :: "actor \ (string set * string set)" defines ex_creds_def: "ex_creds \ (\ x. if x = Actor ''Patient'' then ({''PIN'',''skey''}, {}) else (if x = Actor ''Doctor'' then ({''PIN''},{}) else ({},{})))" fixes ex_locs :: "location \ string * (dlm * data) set" defines "ex_locs \ (\ x. if x = cloud then (''free'',{((Actor ''Patient'',{Actor ''Doctor''}),42)}) else ('''',{}))" fixes ex_loc_ass :: "location \ identity set" defines ex_loc_ass_def: "ex_loc_ass \ (\ x. if x = home then {''Patient''} else (if x = hospital then {''Doctor'', ''Eve''} else {}))" (* The nicer representation with case suffers from not so nice presentation in the cases (need to unfold the syntax) fixes ex_loc_ass_alt :: "location \ identity set" defines ex_loc_ass_alt_def: "ex_loc_ass_alt \ (\ x. (case x of Location (Suc 0) \ {''Patient''} | Location (Suc (Suc 0)) \ {''Doctor'', ''Eve''} | _ \ {}))" *) fixes ex_graph :: "igraph" defines ex_graph_def: "ex_graph \ Lgraph {(home, cloud), (sphone, cloud), (cloud,hospital)} ex_loc_ass ex_creds ex_locs" fixes ex_graph' :: "igraph" defines ex_graph'_def: "ex_graph' \ Lgraph {(home, cloud), (sphone, cloud), (cloud,hospital)} (\ x. if x = cloud then {''Patient''} else (if x = hospital then {''Doctor'',''Eve''} else {})) ex_creds ex_locs" fixes ex_graph'' :: "igraph" defines ex_graph''_def: "ex_graph'' \ Lgraph {(home, cloud), (sphone, cloud), (cloud,hospital)} (\ x. if x = cloud then {''Patient'', ''Eve''} else (if x = hospital then {''Doctor''} else {})) ex_creds ex_locs" (* Same as above: the nicer representation with case suffers from not so nice presentation in the cases (need to unfold the syntax) fixes local_policies_alt :: "[igraph, location] \ policy set" defines local_policies_alt_def: "local_policies_alt G \ (\ x. case x of Location (Suc 0) \ {(\ y. True, {put,get,move,eval})} | Location 0 \ {((\ y. has G (y, ''PIN'')), {put,get,move,eval})} | Location (Suc (Suc (Suc 0))) \ {(\ y. True, {put,get,move,eval})} | Location (Suc (Suc 0)) \ {((\ y. (\ n. (n @\<^bsub>G\<^esub> hospital) \ Actor n = y \ has G (y, ''skey''))), {put,get,move,eval})} | _ \ {})" *) fixes local_policies :: "[igraph, location] \ policy set" defines local_policies_def: "local_policies G \ (\ x. if x = home then {(\ y. True, {put,get,move,eval})} else (if x = sphone then {((\ y. has G (y, ''PIN'')), {put,get,move,eval})} else (if x = cloud then {(\ y. True, {put,get,move,eval})} else (if x = hospital then {((\ y. (\ n. (n @\<^bsub>G\<^esub> hospital) \ Actor n = y \ has G (y, ''skey''))), {put,get,move,eval})} else {}))))" (* problems with case in locales? defines local_policies_def: "local_policies G x \ (case x of home \ {(\ y. True, {put,get,move,eval})} | sphone \ {((\ y. has G (y, ''PIN'')), {put,get,move,eval})} | cloud \ {(\ y. True, {put,get,move,eval})} | hospital \ {((\ y. (\ n. (n @\<^bsub>G\<^esub> hospital) \ Actor n = y \ has G (y, ''skey''))), {put,get,move,eval})} | _ \ {})" *) fixes gdpr_scenario :: "infrastructure" defines gdpr_scenario_def: "gdpr_scenario \ Infrastructure ex_graph local_policies" fixes Igdpr :: "infrastructure set" defines Igdpr_def: "Igdpr \ {gdpr_scenario}" (* other states of scenario *) (* First step: Patient goes onto cloud *) fixes gdpr_scenario' :: "infrastructure" defines gdpr_scenario'_def: "gdpr_scenario' \ Infrastructure ex_graph' local_policies" fixes GDPR' :: "infrastructure set" defines GDPR'_def: "GDPR' \ {gdpr_scenario'}" (* Second step: Eve goes onto cloud from where she'll be able to get the data *) fixes gdpr_scenario'' :: "infrastructure" defines gdpr_scenario''_def: "gdpr_scenario'' \ Infrastructure ex_graph'' local_policies" fixes GDPR'' :: "infrastructure set" defines GDPR''_def: "GDPR'' \ {gdpr_scenario''}" fixes gdpr_states defines gdpr_states_def: "gdpr_states \ { I. gdpr_scenario \\<^sub>i* I }" fixes gdpr_Kripke defines "gdpr_Kripke \ Kripke gdpr_states {gdpr_scenario}" fixes sgdpr defines "sgdpr \ {x. \ (global_policy' x ''Eve'')}" begin subsection \Using Attack Tree Calculus\ text \Since we consider a predicate transformer semantics, we use sets of states to represent properties. For example, the attack property is given by the above @{text \set sgdpr\}. The attack we are interested in is to see whether for the scenario @{text \gdpr scenario \ Infrastructure ex_graph local_policies\} from the initial state @{text \Igdpr \{gdpr scenario}\}, the critical state @{text \sgdpr\} can be reached, i.e., is there a valid attack @{text \(Igdpr,sgdpr)\}? We first present a number of lemmas showing single and multi-step state transitions for relevant states reachable from our @{text \gdpr_scenario\}.\ lemma step1: "gdpr_scenario \\<^sub>n gdpr_scenario'" proof (rule_tac l = home and a = "''Patient''" and l' = cloud in move) show "graphI gdpr_scenario = graphI gdpr_scenario" by (rule refl) next show "''Patient'' @\<^bsub>graphI gdpr_scenario\<^esub> home" by (simp add: gdpr_scenario_def ex_graph_def ex_loc_ass_def atI_def nodes_def) next show "home \ nodes (graphI gdpr_scenario)" by (simp add: gdpr_scenario_def ex_graph_def ex_loc_ass_def atI_def nodes_def, blast) next show "cloud \ nodes (graphI gdpr_scenario)" by (simp add: gdpr_scenario_def nodes_def ex_graph_def, blast) next show "''Patient'' \ actors_graph (graphI gdpr_scenario)" by (simp add: actors_graph_def gdpr_scenario_def ex_graph_def ex_loc_ass_def nodes_def, blast) next show "enables gdpr_scenario cloud (Actor ''Patient'') move" by (simp add: enables_def gdpr_scenario_def ex_graph_def local_policies_def ex_creds_def ex_locs_def has_def credentials_def) next show "gdpr_scenario' = Infrastructure (move_graph_a ''Patient'' home cloud (graphI gdpr_scenario)) (delta gdpr_scenario)" apply (simp add: gdpr_scenario'_def ex_graph'_def move_graph_a_def gdpr_scenario_def ex_graph_def home_def cloud_def hospital_def ex_loc_ass_def ex_creds_def) apply (rule ext) by (simp add: hospital_def) qed lemma step1r: "gdpr_scenario \\<^sub>n* gdpr_scenario'" proof (simp add: state_transition_in_refl_def) show " (gdpr_scenario, gdpr_scenario') \ {(x::infrastructure, y::infrastructure). x \\<^sub>n y}\<^sup>*" by (insert step1, auto) qed lemma step2: "gdpr_scenario' \\<^sub>n gdpr_scenario''" proof (rule_tac l = hospital and a = "''Eve''" and l' = cloud in move, rule refl) show "''Eve'' @\<^bsub>graphI gdpr_scenario'\<^esub> hospital" by (simp add: gdpr_scenario'_def ex_graph'_def hospital_def cloud_def atI_def nodes_def) next show "hospital \ nodes (graphI gdpr_scenario')" by (simp add: gdpr_scenario'_def ex_graph'_def hospital_def cloud_def atI_def nodes_def, blast) next show "cloud \ nodes (graphI gdpr_scenario')" by (simp add: gdpr_scenario'_def nodes_def ex_graph'_def, blast) next show "''Eve'' \ actors_graph (graphI gdpr_scenario')" by (simp add: actors_graph_def gdpr_scenario'_def ex_graph'_def nodes_def hospital_def cloud_def, blast) next show "enables gdpr_scenario' cloud (Actor ''Eve'') move" by (simp add: enables_def gdpr_scenario'_def ex_graph_def local_policies_def ex_creds_def ex_locs_def has_def credentials_def cloud_def sphone_def) next show "gdpr_scenario'' = Infrastructure (move_graph_a ''Eve'' hospital cloud (graphI gdpr_scenario')) (delta gdpr_scenario')" apply (simp add: gdpr_scenario'_def ex_graph''_def move_graph_a_def gdpr_scenario''_def ex_graph'_def home_def cloud_def hospital_def ex_creds_def) apply (rule ext) apply (simp add: hospital_def) by blast qed lemma step2r: "gdpr_scenario' \\<^sub>n* gdpr_scenario''" proof (simp add: state_transition_in_refl_def) show "(gdpr_scenario', gdpr_scenario'') \ {(x::infrastructure, y::infrastructure). x \\<^sub>n y}\<^sup>*" by (insert step2, auto) qed (* Attack example: Eve can get onto cloud and get Patient's data because the policy allows Eve to get on cloud. This attack can easily be fixed by disabling Eve to "get" in the policy (just change the "True" for cloud to a set with no Eve in it). However, it would not prevent Insider attacks (where Eve is impersonating the Doctor, for example). Insider attacks can be checked using the UasI predicate. *) text \For the Kripke structure @{text \gdpr_Kripke \ Kripke { I. gdpr_scenario \\<^sub>i* I } {gdpr_scenario}\} we first derive a valid and-attack using the attack tree proof calculus. @{text \"\[\\<^bsub>(Igdpr,GDPR')\<^esub>, \\<^bsub>(GDPR',sgdpr)\<^esub>] \\<^sub>\\<^bsup>(Igdpr,sgdpr)\<^esup>\} The set @{text \GDPR'\} (see above) is an intermediate state where Eve accesses the cloud.\ lemma gdpr_ref: "[\\<^bsub>(Igdpr,sgdpr)\<^esub>] \\<^sub>\\<^bsup>(Igdpr,sgdpr)\<^esup> \ ([\\<^bsub>(Igdpr,GDPR')\<^esub>, \\<^bsub>(GDPR',sgdpr)\<^esub>] \\<^sub>\\<^bsup>(Igdpr,sgdpr)\<^esup>)" proof (rule_tac l = "[]" and l' = "[\\<^bsub>(Igdpr,GDPR')\<^esub>, \\<^bsub>(GDPR',sgdpr)\<^esub>]" and l'' = "[]" and si = Igdpr and si' = Igdpr and si'' = sgdpr and si''' = sgdpr in refI, simp, rule refl) show "([\\<^bsub>(Igdpr, GDPR')\<^esub>, \\<^bsub>(GDPR', sgdpr)\<^esub>] \\<^sub>\\<^bsup>(Igdpr, sgdpr)\<^esup>) = ([] @ [\\<^bsub>(Igdpr, GDPR')\<^esub>, \\<^bsub>(GDPR', sgdpr)\<^esub>] @ [] \\<^sub>\\<^bsup>(Igdpr, sgdpr)\<^esup>)" by simp qed lemma att_gdpr: "\([\\<^bsub>(Igdpr,GDPR')\<^esub>, \\<^bsub>(GDPR',sgdpr)\<^esub>] \\<^sub>\\<^bsup>(Igdpr,sgdpr)\<^esup>)" proof (subst att_and, simp, rule conjI) show " \\\<^bsub>(Igdpr, GDPR')\<^esub>" apply (simp add: Igdpr_def GDPR'_def att_base) using state_transition_infra_def step1 by blast next have "\ global_policy' gdpr_scenario'' ''Eve''" "gdpr_scenario' \\<^sub>n gdpr_scenario''" using step2 by (auto simp: global_policy'_def gdpr_scenario''_def gdpr_actors_def enables_def local_policies_def cloud_def sphone_def intro!: step2) then show "\([\\<^bsub>(GDPR', sgdpr)\<^esub>] \\<^sub>\\<^bsup>(GDPR', sgdpr)\<^esup>)" apply (subst att_and) apply (simp add: GDPR'_def sgdpr_def att_base) using state_transition_infra_def by blast qed lemma gdpr_abs_att: "\\<^sub>V([\\<^bsub>(Igdpr,sgdpr)\<^esub>] \\<^sub>\\<^bsup>(Igdpr,sgdpr)\<^esup>)" by (rule ref_valI, rule gdpr_ref, rule att_gdpr) text \We can then simply apply the Correctness theorem @{text \AT EF\} to immediately prove the following CTL statement. @{text \gdpr_Kripke \ EF sgdpr\} This application of the meta-theorem of Correctness of attack trees saves us proving the CTL formula tediously by exploring the state space.\ lemma gdpr_att: "gdpr_Kripke \ EF {x. \(global_policy' x ''Eve'')}" proof - have a: " \([\\<^bsub>(Igdpr, GDPR')\<^esub>, \\<^bsub>(GDPR', sgdpr)\<^esub>] \\<^sub>\\<^bsup>(Igdpr, sgdpr)\<^esup>)" by (rule att_gdpr) hence "(Igdpr,sgdpr) = attack ([\\<^bsub>(Igdpr, GDPR')\<^esub>, \\<^bsub>(GDPR', sgdpr)\<^esub>] \\<^sub>\\<^bsup>(Igdpr, sgdpr)\<^esup>)" by simp hence "Kripke {s::infrastructure. \i::infrastructure\Igdpr. i \\<^sub>i* s} Igdpr \ EF sgdpr" using ATV_EF gdpr_abs_att by fastforce thus "gdpr_Kripke \ EF {x::infrastructure. \ global_policy' x ''Eve''}" by (simp add: gdpr_Kripke_def gdpr_states_def Igdpr_def sgdpr_def) qed theorem gdpr_EF: "gdpr_Kripke \ EF sgdpr" using gdpr_att sgdpr_def by blast text \Similarly, vice-versa, the CTL statement proved in @{text \gdpr_EF\} can now be directly translated into Attack Trees using the Completeness Theorem\footnote{This theorem could easily be proved as a direct instance of @{text \att_gdpr\} above but we want to illustrate an alternative proof method using Completeness here.}.\ theorem gdpr_AT: "\ A. \ A \ attack A = (Igdpr,sgdpr)" proof - have a: "gdpr_Kripke \ EF sgdpr " by (rule gdpr_EF) have b: "Igdpr \ {}" by (simp add: Igdpr_def) thus "\A::infrastructure attree. \A \ attack A = (Igdpr, sgdpr)" proof (rule Completeness) show "Kripke {s. \i\Igdpr. i \\<^sub>i* s} Igdpr \ EF sgdpr" using a by (simp add: gdpr_Kripke_def Igdpr_def gdpr_states_def) qed (auto simp: Igdpr_def) qed text \Conversely, since we have an attack given by rule @{text \gdpr_AT\}, we can immediately infer @{text \EF s\} using Correctness @{text \AT_EF\}\footnote{Clearly, this theorem is identical to @{text \gdpr_EF\} and could thus be inferred from that one but we want to show here an alternative way of proving it using the Correctness theorem @{text \AT_EF\}.}.\ theorem gdpr_EF': "gdpr_Kripke \ EF sgdpr" using gdpr_AT by (auto simp: gdpr_Kripke_def gdpr_states_def Igdpr_def dest: AT_EF) (* However, when integrating DLM into the model and hence labeling information becomes part of the conditions of the get_data rule this isn't possible any more: gdpr_EF is not true any more *) (** GDPR properties for the illustration of the DLM labeling **) section \Data Protection by Design for GDPR compliance\ subsection \General Data Protection Regulation (GDPR)\ text \Since 26th May 2018, the GDPR has become mandatory within the European Union and hence also for any supplier of IT products. Breaches of the regulation will be fined with penalties of 20 Million EUR. Despite the relatively large size of the document of 209 pages, the technically relevant portion for us is only about 30 pages (Pages 81--111, Chapters I to Chapter III, Section 3). In summary, Chapter III specifies that the controller must give the data subject read access (1) -to any information, communications, and “meta-data” of the data, e.g., retention time and purpose. +to any information, communications, and ``meta-data'' of the data, e.g., retention time and purpose. In addition, the system must enable deletion of data (2) and restriction of processing. An invariant condition for data processing resulting from these Articles is that the system functions must preserve any of the access rights of personal data (3). Using labeled data, we can now express the essence of Article 4 Paragraph (1): -’personal data’ means any information relating to an identified or identifiable natural -person (’data subject’). +`personal data' means any information relating to an identified or identifiable natural +person (`data subject'). The labels of data must not be changed by processing: we have identified this as an invariant (3) resulting from the GDPR above. This invariant is formalized in our Isabelle model by the type definition of functions on labeled data @{text \label_fun\} (see Section 4.2) that preserve the data labels.\ subsection \Policy enforcement and privacy preservation\ text \We can now use the labeled data to encode the privacy constraints of the GDPR in the rules. For example, the get data rule (see Section 4.3) has labelled data - @{text \((Actor a’, as), n)\} and uses the labeling in the precondition to guarantee + @{text \((Actor a', as), n)\} and uses the labeling in the precondition to guarantee that only entitled users can get data. We can prove that processing preserves ownership as defined in the initial state for all paths globally (AG) within the Kripke structure and in all locations of the graph.\ (* GDPR three: Processing preserves ownership in any location *) lemma gdpr_three: "h \ gdpr_actors \ l \ gdpr_locations \ owns (Igraph gdpr_scenario) l (Actor h) d \ gdpr_Kripke \ AG {x. \ l \ gdpr_locations. owns (Igraph x) l (Actor h) d }" proof (simp add: gdpr_Kripke_def check_def, rule conjI) show "gdpr_scenario \ gdpr_states" by (simp add: gdpr_states_def state_transition_refl_def) next show "h \ gdpr_actors \ l \ gdpr_locations \ owns (Igraph gdpr_scenario) l (Actor h) d \ gdpr_scenario \ AG {x::infrastructure. \l\gdpr_locations. owns (Igraph x) l (Actor h) d}" apply (simp add: AG_def gfp_def) apply (rule_tac x = "{x::infrastructure. \l\gdpr_locations. owns (Igraph x) l (Actor h) d}" in exI) by (auto simp: AX_def gdpr_scenario_def owns_def) qed text \The final application example of Correctness contraposition shows that there is no attack to ownership possible. The proved meta-theory for attack trees can be applied to facilitate the proof. The contraposition of the Correctness property grants that if there is no attack on @{text \(I,\f)\}, then @{text \(EF \f)\} does not hold in the Kripke structure. This yields the theorem since the @{text \AG f\} statement corresponds to @{text \\(EF \f)\}. \ theorem no_attack_gdpr_three: "h \ gdpr_actors \ l \ gdpr_locations \ owns (Igraph gdpr_scenario) l (Actor h) d \ attack A = (Igdpr, -{x. \ l \ gdpr_locations. owns (Igraph x) l (Actor h) d }) \ \ (\ A)" proof (rule_tac I = Igdpr and s = "-{x::infrastructure. \l\gdpr_locations. owns (Igraph x) l (Actor h) d}" in contrapos_corr) show "h \ gdpr_actors \ l \ gdpr_locations \ owns (Igraph gdpr_scenario) l (Actor h) d \ attack A = (Igdpr, - {x::infrastructure. \l\gdpr_locations. owns (Igraph x) l (Actor h) d}) \ \ (Kripke {s::infrastructure. \i::infrastructure\Igdpr. i \\<^sub>i* s} Igdpr \ EF (- {x::infrastructure. \l\gdpr_locations. owns (Igraph x) l (Actor h) d}))" apply (rule AG_imp_notnotEF) apply (simp add: gdpr_Kripke_def Igdpr_def gdpr_states_def) using Igdpr_def gdpr_Kripke_def gdpr_states_def gdpr_three by auto qed end end \ No newline at end of file diff --git a/thys/Attack_Trees/MC.thy b/thys/Attack_Trees/MC.thy --- a/thys/Attack_Trees/MC.thy +++ b/thys/Attack_Trees/MC.thy @@ -1,467 +1,467 @@ section "Kripke structures and CTL" text \We apply Kripke structures and CTL to model state based systems and analyse properties under dynamic state changes. Snapshots of systems are the states on which we define a state transition. Temporal logic is then employed to express security and privacy properties.\ theory MC imports Main begin subsection "Lemmas to support least and greatest fixpoints" lemma predtrans_empty: assumes "mono (\ :: 'a set \ 'a set)" shows "\ i. (\ ^^ i) ({}) \ (\ ^^(i + 1))({})" using assms funpow_decreasing le_add1 by blast lemma ex_card: "finite S \ \ n:: nat. card S = n" by simp lemma less_not_le: "\(x:: nat) < y; y \ x\ \ False" by arith lemma infchain_outruns_all: assumes "finite (UNIV :: 'a set)" and "\i :: nat. ((\ :: 'a set \ 'a set)^^ i) ({}:: 'a set) \ (\ ^^ (i + 1)) {}" shows "\j :: nat. \i :: nat. j < card ((\ ^^ i) {})" proof (rule allI, induct_tac j) show "\i. 0 < card ((\ ^^ i) {})" using assms by (metis bot.not_eq_extremum card_gt_0_iff finite_subset subset_UNIV) next show "\j n. \i. n < card ((\ ^^ i) {}) \ \i. Suc n < card ((\ ^^ i) {})" proof - fix j n assume a: "\i. n < card ((\ ^^ i) {})" obtain i where "n < card ((\ ^^ i) {})" using a by blast thus "\ i. Suc n < card ((\ ^^ i) {})" using assms by (meson finite_subset le_less_trans le_simps(3) psubset_card_mono subset_UNIV) qed qed lemma no_infinite_subset_chain: assumes "finite (UNIV :: 'a set)" and "mono (\ :: ('a set \ 'a set))" and "\i :: nat. ((\ :: 'a set \ 'a set) ^^ i) {} \ (\ ^^ (i + (1 :: nat))) ({} :: 'a set)" shows "False" text \Proof idea: since @{term "UNIV"} is finite, we have from @{text \ex_card\} that there is an n with @{term "card UNIV = n"}. Now, use @{text \infchain_outruns_all\} to show as contradiction point that @{term "\ i :: nat. card UNIV < card ((\ ^^ i) {})"}. Since all sets are subsets of @{term "UNIV"}, we also have @{term "card ((\ ^^ i) {}) \ card UNIV"}: Contradiction!, i.e. proof of False \ proof - have a: "\ (j :: nat). (\ (i :: nat). (j :: nat) < card((\ ^^ i)({} :: 'a set)))" using assms by (erule_tac \ = \ in infchain_outruns_all) hence b: "\ (n :: nat). card(UNIV :: 'a set) = n" using assms by (erule_tac S = UNIV in ex_card) from this obtain n where c: "card(UNIV :: 'a set) = n" by (erule exE) hence d: "\i. card UNIV < card ((\ ^^ i) {})" using a by (drule_tac x = "card UNIV" in spec) from this obtain i where e: "card (UNIV :: 'a set) < card ((\ ^^ i) {})" by (erule exE) hence f: "(card((\ ^^ i){})) \ (card (UNIV :: 'a set))" using assms apply (erule_tac A = "((\ ^^ i){})" in Finite_Set.card_mono) by (rule subset_UNIV) thus "False" using e by (erule_tac y = "card((\ ^^ i){})" in less_not_le) qed lemma finite_fixp: assumes "finite(UNIV :: 'a set)" and "mono (\ :: ('a set \ 'a set))" shows "\ i. (\ ^^ i) ({}) = (\ ^^(i + 1))({})" text \Proof idea: with @{text predtrans_empty} we know @{term "\ i. (\ ^^ i){} \ (\ ^^(i + 1))({})"} (1). If we can additionally show @{term "\ i. (\ ^^ i)({}) \ (\ ^^(i + 1))({})"} (2), we can get the goal together with equalityI @{text "\ + \ \ ="}. To prove (1) we observe that @{term "(\ ^^ i)({}) \ (\ ^^(i + 1))({})"} can be inferred from @{term "\((\ ^^ i)({}) \ (\ ^^(i + 1))({}))"} and (1). Finally, the latter is solved directly by @{text \no_infinite_subset_chain\}.\ proof - have a: "\i. (\ ^^ i) ({}:: 'a set) \ (\ ^^ (i + (1))) {}" by(rule predtrans_empty, rule assms(2)) have a3: "\ (\ i :: nat. (\ ^^ i) {} \ (\ ^^(i + 1)) {})" by (rule notI, rule no_infinite_subset_chain, (rule assms)+) hence b: "(\ i :: nat. \((\ ^^ i) {} \ (\ ^^(i + 1)) {}))" using assms a3 by blast thus "\ i. (\ ^^ i) ({}) = (\ ^^(i + 1))({})" using a by blast qed lemma predtrans_UNIV: assumes "mono (\ :: ('a set \ 'a set))" shows "\ i. (\ ^^ i) (UNIV) \ (\ ^^(i + 1))(UNIV)" proof (rule allI, induct_tac i) show "(\ ^^ ((0) + (1))) UNIV \ (\ ^^ (0)) UNIV" by simp next show "\(i) n. (\ ^^ (n + (1))) UNIV \ (\ ^^ n) UNIV \ (\ ^^ (Suc n + (1))) UNIV \ (\ ^^ Suc n) UNIV" proof - fix i n assume a: "(\ ^^ (n + (1))) UNIV \ (\ ^^ n) UNIV" have "(\ ((\ ^^ n) UNIV)) \ (\ ((\ ^^ (n + (1 :: nat))) UNIV))" using assms a by (rule monoE) thus "(\ ^^ (Suc n + (1))) UNIV \ (\ ^^ Suc n) UNIV" by simp qed qed lemma Suc_less_le: "x < (y - n) \ x \ (y - (Suc n))" by simp lemma card_univ_subtract: assumes "finite (UNIV :: 'a set)" and "mono \" and "(\i :: nat. ((\ :: 'a set \ 'a set) ^^ (i + (1 :: nat)))(UNIV :: 'a set) \ (\ ^^ i) UNIV)" shows "(\ i :: nat. card((\ ^^ i) (UNIV ::'a set)) \ (card (UNIV :: 'a set)) - i)" proof (rule allI, induct_tac i) show "card ((\ ^^ (0)) UNIV) \ card (UNIV :: 'a set) - (0)" using assms by (simp) next show "\(i) n. card ((\ ^^ n) (UNIV :: 'a set)) \ card (UNIV :: 'a set) - n \ card ((\ ^^ Suc n) (UNIV :: 'a set)) \ card (UNIV :: 'a set) - Suc n" using assms proof - fix i n assume a: "card ((\ ^^ n) (UNIV :: 'a set)) \ card (UNIV :: 'a set) - n" have b: "(\ ^^ (n + (1)))(UNIV :: 'a set) \ (\ ^^ n) UNIV" using assms by (erule_tac x = n in spec) have "card((\ ^^ (n + (1 :: nat)))(UNIV :: 'a set)) < card((\ ^^ n) (UNIV :: 'a set))" by (rule psubset_card_mono, rule finite_subset, rule subset_UNIV, rule assms(1), rule b) thus "card ((\ ^^ Suc n) (UNIV :: 'a set)) \ card (UNIV :: 'a set) - Suc n" using a by simp qed qed lemma card_UNIV_tau_i_below_zero: assumes "finite (UNIV :: 'a set)" and "mono \" and "(\i :: nat. ((\ :: ('a set \ 'a set)) ^^ (i + (1 :: nat)))(UNIV :: 'a set) \ (\ ^^ i) UNIV)" shows "card((\ ^^ (card (UNIV ::'a set))) (UNIV ::'a set)) \ 0" proof - have "(\ i :: nat. card((\ ^^ i) (UNIV ::'a set)) \ (card (UNIV :: 'a set)) - i)" using assms by (rule card_univ_subtract) thus "card((\ ^^ (card (UNIV ::'a set))) (UNIV ::'a set)) \ 0" by (drule_tac x = "card (UNIV ::'a set)" in spec, simp) qed lemma finite_card_zero_empty: "\ finite S; card S \ 0\ \ S = {}" by simp lemma UNIV_tau_i_is_empty: assumes "finite (UNIV :: 'a set)" and "mono (\ :: ('a set \ 'a set))" and "(\i :: nat. ((\ :: 'a set \ 'a set) ^^ (i + (1 :: nat)))(UNIV :: 'a set) \ (\ ^^ i) UNIV)" shows "(\ ^^ (card (UNIV ::'a set))) (UNIV ::'a set) = {}" by (meson assms card_UNIV_tau_i_below_zero finite_card_zero_empty finite_subset subset_UNIV) lemma down_chain_reaches_empty: assumes "finite (UNIV :: 'a set)" and "mono (\ :: 'a set \ 'a set)" and "(\i :: nat. ((\ :: 'a set \ 'a set) ^^ (i + (1 :: nat))) UNIV \ (\ ^^ i) UNIV)" shows "\ (j :: nat). (\ ^^ j) UNIV = {}" using UNIV_tau_i_is_empty assms by blast lemma no_infinite_subset_chain2: assumes "finite (UNIV :: 'a set)" and "mono (\ :: ('a set \ 'a set))" and "\i :: nat. (\ ^^ i) UNIV \ (\ ^^ (i + (1 :: nat))) UNIV" shows "False" proof - have "\ j :: nat. (\ ^^ j) UNIV = {}" using assms by (rule down_chain_reaches_empty) from this obtain j where a: "(\ ^^ j) UNIV = {}" by (erule exE) have "(\ ^^ (j + (1))) UNIV \ (\ ^^ j) UNIV" using assms by (erule_tac x = j in spec) thus False using a by simp qed lemma finite_fixp2: assumes "finite(UNIV :: 'a set)" and "mono (\ :: ('a set \ 'a set))" shows "\ i. (\ ^^ i) UNIV = (\ ^^(i + 1)) UNIV" proof - have "\i. (\ ^^ (i + (1))) UNIV \ (\ ^^ i) UNIV" by (rule predtrans_UNIV , simp add: assms(2)) moreover have "\i. \ (\ ^^ (i + (1))) UNIV \ (\ ^^ i) UNIV" using assms proof - have "\ (\ i :: nat. (\ ^^ i) UNIV \ (\ ^^(i + 1)) UNIV)" using assms(1) assms(2) no_infinite_subset_chain2 by blast thus "\i. \ (\ ^^ (i + (1))) UNIV \ (\ ^^ i) UNIV" by blast qed ultimately show "\ i. (\ ^^ i) UNIV = (\ ^^(i + 1)) UNIV" by blast qed lemma lfp_loop: assumes "finite (UNIV :: 'b set)" and "mono (\ :: ('b set \ 'b set))" shows "\ n . lfp \ = (\ ^^ n) {}" proof - have "\i. (\ ^^ i) {} = (\ ^^ (i + (1))) {}" using assms by (rule finite_fixp) from this obtain i where " (\ ^^ i) {} = (\ ^^ (i + (1))) {}" by (erule exE) hence "(\ ^^ i) {} = (\ ^^ Suc i) {}" by simp hence "(\ ^^ Suc i) {} = (\ ^^ i) {}" by (rule sym) hence "lfp \ = (\ ^^ i) {}" by (simp add: assms(2) lfp_Kleene_iter) thus "\ n . lfp \ = (\ ^^ n) {}" by (rule exI) qed text \These next two are repeated from the corresponding theorems in HOL/ZF/Nat.thy for the sake of self-containedness of the exposition.\ lemma Kleene_iter_gpfp: assumes "mono f" and "p \ f p" shows "p \ (f^^k) (top::'a::order_top)" proof(induction k) case 0 show ?case by simp next case Suc from monoD[OF assms(1) Suc] assms(2) show ?case by simp qed lemma gfp_loop: assumes "finite (UNIV :: 'b set)" and "mono (\ :: ('b set \ 'b set))" shows "\ n . gfp \ = (\ ^^ n)UNIV" proof - have " \i. (\ ^^ i)(UNIV :: 'b set) = (\ ^^ (i + (1))) UNIV" using assms by (rule finite_fixp2) from this obtain i where "(\ ^^ i)UNIV = (\ ^^ (i + (1))) UNIV" by (erule exE) thus "\ n . gfp \ = (\ ^^ n)UNIV" using assms by (metis Suc_eq_plus1 gfp_Kleene_iter) qed subsection \Generic type of state with state transition and CTL operators\ text \The system states and their transition relation are defined as a class called @{text \state\} containing an abstract constant@{text \state_transition\}. It introduces the -syntactic infix notation @{text \I \\<^sub>i I'\} to denote that system state @{text \I\} and @{text \I’\} +syntactic infix notation @{text \I \\<^sub>i I'\} to denote that system state @{text \I\} and @{text \I'\} are in this relation over an arbitrary (polymorphic) type @{text \'a\}.\ class state = fixes state_transition :: "['a :: type, 'a] \ bool" (infixr "\\<^sub>i" 50) text \The above class definition lifts Kripke structures and CTL to a general level. The definition of the inductive relation is given by a set of specific rules which are, however, part of an application like infrastructures. Branching time temporal logic CTL is defined in general over Kripke structures with arbitrary state transitions and can later be applied to suitable theories, like infrastructures. Based on the generic state transition @{text \\\} of the type class state, the CTL-operators EX and AX express that property f holds in some or all next states, respectively.\ definition AX where "AX f \ {s. {f0. s \\<^sub>i f0} \ f}" definition EX' where "EX' f \ {s . \ f0 \ f. s \\<^sub>i f0 }" text \The CTL formula @{text \AG f\} means that on all paths branching from a state @{text \s\} -the formula @{text \f\} is always true (@{text \G\} stands for ‘globally’). It can be defined +the formula @{text \f\} is always true (@{text \G\} stands for `globally'). It can be defined using the Tarski fixpoint theory by applying the greatest fixpoint operator. In a similar way, the other CTL operators are defined.\ definition AF where "AF f \ lfp (\ Z. f \ AX Z)" definition EF where "EF f \ lfp (\ Z. f \ EX' Z)" definition AG where "AG f \ gfp (\ Z. f \ AX Z)" definition EG where "EG f \ gfp (\ Z. f \ EX' Z)" definition AU where "AU f1 f2 \ lfp(\ Z. f2 \ (f1 \ AX Z))" definition EU where "EU f1 f2 \ lfp(\ Z. f2 \ (f1 \ EX' Z))" definition AR where "AR f1 f2 \ gfp(\ Z. f2 \ (f1 \ AX Z))" definition ER where "ER f1 f2 \ gfp(\ Z. f2 \ (f1 \ EX' Z))" subsection \Kripke structures and Modelchecking\ datatype 'a kripke = Kripke "'a set" "'a set" primrec states where "states (Kripke S I) = S" primrec init where "init (Kripke S I) = I" text \The formal Isabelle definition of what it means that formula f holds in a Kripke structure M can be stated as: the initial states of the Kripke structure init M need to be contained in the set of all states states M that imply f.\ definition check ("_ \ _" 50) where "M \ f \ (init M) \ {s \ (states M). s \ f }" definition state_transition_refl (infixr "\\<^sub>i*" 50) where "s \\<^sub>i* s' \ ((s,s') \ {(x,y). state_transition x y}\<^sup>*)" subsection \Lemmas for CTL operators\ subsubsection \EF lemmas\ lemma EF_lem0: "(x \ EF f) = (x \ f \ EX' (lfp (\Z :: ('a :: state) set. f \ EX' Z)))" proof - have "lfp (\Z :: ('a :: state) set. f \ EX' Z) = f \ (EX' (lfp (\Z :: 'a set. f \ EX' Z)))" by (rule def_lfp_unfold, rule reflexive, unfold mono_def EX'_def, auto) thus "(x \ EF (f :: ('a :: state) set)) = (x \ f \ EX' (lfp (\Z :: ('a :: state) set. f \ EX' Z)))" by (simp add: EF_def) qed lemma EF_lem00: "(EF f) = (f \ EX' (lfp (\ Z :: ('a :: state) set. f \ EX' Z)))" by (auto simp: EF_lem0) lemma EF_lem000: "(EF f) = (f \ EX' (EF f))" by (metis EF_def EF_lem00) lemma EF_lem1: "x \ f \ x \ (EX' (EF f)) \ x \ EF f" proof (simp add: EF_def) assume a: "x \ f \ x \ EX' (lfp (\Z::'a set. f \ EX' Z))" show "x \ lfp (\Z::'a set. f \ EX' Z)" proof - have b: "lfp (\Z :: ('a :: state) set. f \ EX' Z) = f \ (EX' (lfp (\Z :: ('a :: state) set. f \ EX' Z)))" using EF_def EF_lem00 by blast thus "x \ lfp (\Z::'a set. f \ EX' Z)" using a by (subst b, blast) qed qed lemma EF_lem2b: assumes "x \ (EX' (EF f))" shows "x \ EF f" by (simp add: EF_lem1 assms) lemma EF_lem2a: assumes "x \ f" shows "x \ EF f" by (simp add: EF_lem1 assms) lemma EF_lem2c: assumes "x \ f" shows "x \ EF (- f)" by (simp add: EF_lem1 assms) lemma EF_lem2d: assumes "x \ EF f" shows "x \ f" using EF_lem1 assms by auto lemma EF_lem3b: assumes "x \ EX' (f \ EX' (EF f))" shows "x \ (EF f)" by (metis EF_lem000 EF_lem2b assms) lemma EX_lem0l: "x \ (EX' f) \ x \ (EX' (f \ g))" by (auto simp: EX'_def) lemma EX_lem0r: "x \ (EX' g) \ x \ (EX' (f \ g))" by (auto simp: EX'_def) lemma EX_step: assumes "x \\<^sub>i y" and "y \ f" shows "x \ EX' f" using assms by (auto simp: EX'_def) lemma EF_E[rule_format]: "\ f. x \ (EF f) \ x \ (f \ EX' (EF f))" using EF_lem000 by blast lemma EF_step: assumes "x \\<^sub>i y" and "y \ f" shows "x \ EF f" using EF_lem3b EX_step assms by blast lemma EF_step_step: assumes "x \\<^sub>i y" and "y \ EF f" shows "x \ EF f" using EF_lem2b EX_step assms by blast lemma EF_step_star: "\ x \\<^sub>i* y; y \ f \ \ x \ EF f" proof (simp add: state_transition_refl_def) show "(x, y) \ {(x::'a, y::'a). x \\<^sub>i y}\<^sup>* \ y \ f \ x \ EF f" proof (erule converse_rtrancl_induct) show "y \ f \ y \ EF f" by (erule EF_lem2a) next show "\ya z::'a. y \ f \ (ya, z) \ {(x,y). x \\<^sub>i y} \ (z, y) \ {(x,y). x \\<^sub>i y}\<^sup>* \ z \ EF f \ ya \ EF f" by (simp add: EF_step_step) qed qed lemma EF_induct: "(a::'a::state) \ EF f \ mono (\ Z. f \ EX' Z) \ (\x. x \ ((\ Z. f \ EX' Z)(EF f \ {x::'a::state. P x})) \ P x) \ P a" by (metis (mono_tags, lifting) EF_def def_lfp_induct_set) lemma valEF_E: "M \ EF f \ x \ init M \ x \ EF f" by (auto simp: check_def) lemma EF_step_star_rev[rule_format]: "x \ EF s \ (\ y \ s. x \\<^sub>i* y)" proof (erule EF_induct) show "mono (\Z::'a set. s \ EX' Z)" by (force simp add: mono_def EX'_def) next show "\x::'a. x \ s \ EX' (EF s \ {x::'a. \y::'a\s. x \\<^sub>i* y}) \ \y::'a\s. x \\<^sub>i* y" apply (erule UnE) using state_transition_refl_def apply auto[1] by (auto simp add: EX'_def state_transition_refl_def intro: converse_rtrancl_into_rtrancl) qed lemma EF_step_inv: "(I \ {sa::'s :: state. (\i\I. i \\<^sub>i* sa) \ sa \ EF s}) \ \ x \ I. \ y \ s. x \\<^sub>i* y" using EF_step_star_rev by fastforce subsubsection \AG lemmas\ lemma AG_in_lem: "x \ AG s \ x \ s" by (auto simp add: AG_def gfp_def) lemma AG_lem1: "x \ s \ x \ (AX (AG s)) \ x \ AG s" proof (simp add: AG_def) have "gfp (\Z::'a set. s \ AX Z) = s \ (AX (gfp (\Z::'a set. s \ AX Z)))" by (rule def_gfp_unfold) (auto simp: mono_def AX_def) then show "x \ s \ x \ AX (gfp (\Z::'a set. s \ AX Z)) \ x \ gfp (\Z::'a set. s \ AX Z)" by blast qed lemma AG_lem2: "x \ AG s \ x \ (s \ (AX (AG s)))" proof - have a: "AG s = s \ (AX (AG s))" unfolding AG_def by (rule def_gfp_unfold) (auto simp: mono_def AX_def) thus "x \ AG s \ x \ (s \ (AX (AG s)))" by (erule subst) qed lemma AG_lem3: "AG s = (s \ (AX (AG s)))" using AG_lem1 AG_lem2 by blast lemma AG_step: "y \\<^sub>i z \ y \ AG s \ z \ AG s" using AG_lem2 AX_def by blast lemma AG_all_s: " x \\<^sub>i* y \ x \ AG s \ y \ AG s" proof (simp add: state_transition_refl_def) show "(x, y) \ {(x,y). x \\<^sub>i y}\<^sup>* \ x \ AG s \ y \ AG s" by (erule rtrancl_induct) (auto simp add: AG_step) qed lemma AG_imp_notnotEF: "I \ {} \ ((Kripke {s. \ i \ I. (i \\<^sub>i* s)} I \ AG s)) \ (\(Kripke {s. \ i \ I. (i \\<^sub>i* s)} (I :: ('s :: state)set) \ EF (- s)))" proof (rule notI, simp add: check_def) assume a0: "I \ {}" and a1: "I \ {sa::'s. (\i\I. i \\<^sub>i* sa) \ sa \ AG s}" and a2: "I \ {sa::'s. (\i\I. i \\<^sub>i* sa) \ sa \ EF (- s)}" show "False" proof - have a3: "{sa::'s. (\i\I. i \\<^sub>i* sa) \ sa \ AG s} \ {sa::'s. (\i\I. i \\<^sub>i* sa) \ sa \ EF (- s)} = {}" proof - have "(? x. x \ {sa::'s. (\i\I. i \\<^sub>i* sa) \ sa \ AG s} \ x \ {sa::'s. (\i\I. i \\<^sub>i* sa) \ sa \ EF (- s)}) \ False" proof - assume a4: "(? x. x \ {sa::'s. (\i\I. i \\<^sub>i* sa) \ sa \ AG s} \ x \ {sa::'s. (\i\I. i \\<^sub>i* sa) \ sa \ EF (- s)})" from a4 obtain x where a5: "x \ {sa::'s. (\i\I. i \\<^sub>i* sa) \ sa \ AG s} \ x \ {sa::'s. (\i\I. i \\<^sub>i* sa) \ sa \ EF (- s)}" by (erule exE) thus "False" by (metis (mono_tags, lifting) AG_all_s AG_in_lem ComplD EF_step_star_rev a5 mem_Collect_eq) qed thus "{sa::'s. (\i\I. i \\<^sub>i* sa) \ sa \ AG s} \ {sa::'s. (\i\I. i \\<^sub>i* sa) \ sa \ EF (- s)} = {}" by blast qed moreover have b: "? x. x : I" using a0 by blast moreover obtain x where "x \ I" using b by blast ultimately show "False" using a0 a1 a2 by blast qed qed text \A simplified way of Modelchecking is given by the following lemma.\ lemma check2_def: "(Kripke S I \ f) = (I \ S \ f)" by (auto simp add: check_def) end \ No newline at end of file diff --git a/thys/Belief_Revision/AGM_Logic.thy b/thys/Belief_Revision/AGM_Logic.thy --- a/thys/Belief_Revision/AGM_Logic.thy +++ b/thys/Belief_Revision/AGM_Logic.thy @@ -1,480 +1,480 @@ (*<*) \\ ******************************************************************** * Project : AGM Theory * Version : 1.0 * * Authors : Valentin Fouillard, Safouan Taha, Frederic Boulanger and Nicolas Sabouret * * This file : AGM logics * * Copyright (c) 2021 Université Paris Saclay, France * ******************************************************************************\ theory AGM_Logic imports Main begin (*>*) section \Introduction\ text\ The 1985 paper by Carlos Alchourrón, Peter Gärdenfors, -and David Makinson (AGM), “On the Logic of Theory Change: Partial Meet -Contraction and Revision Functions” @{cite "alchourron1985logic"} launches a large and +and David Makinson (AGM), ``On the Logic of Theory Change: Partial Meet +Contraction and Revision Functions'' @{cite "alchourron1985logic"} launches a large and rapidly growing literature that employs formal models and logics to handle changing beliefs of a rational agent and to take into account new piece of information observed by this agent. In 2011, a review book titled "AGM 25 Years: Twenty-Five Years of Research in Belief Change" was edited to summarize the first twenty five years of works based on AGM @{cite "Ferme2011"}. According to Google Scholar, the original AGM paper was cited 4000 times! This AFP entry is HOL-based and it is a faithful formalization of the logic operators (e.g. contraction, revision, remainder \dots ) axiomatized in the AGM paper. It also contains the proofs of all the theorems stated in the paper that show how these operators combine. Both proofs of Harper and Levi identities are established. \ text\A belief state can be considered as a consistent set of beliefs (logical propositions) closed under logical reasoning. Belief changes represent the operations that apply on a belief state to remove some of it and/or to add new beliefs (propositions). In the latter case, it is possible that other beliefs are affected by these changes (to preserve consistency for example). AGM define several postulates to guarantee that such operations preserve consistency meaning that the agent keeps rational. Three kinds of operators are defined : \<^item> The contraction @{text \\
\} : where a proposition is removed from a belief set \<^item> The expansion @{text \\\} : where a proposition is added to a belief set \<^item> The revision @{text \\<^bold>*\} : where a proposition is added to a belief set such that the belief set remains consistent \ text\In this AFP entry, there are three theory files: \<^enum> The AGM Logic file contains a classification of logics used in the AGM framework. \<^enum> The AGM Remainder defines a important operator used in the AGM framework. \<^enum> The AGM Contraction file contains the postulates of the AGM contraction and its relation with the meet contraction. \<^enum> The AGM Revision file contains the postulates of the AGM revision and its relation with the meet revision. \ section \Logics\ text\The AGM framework depends on the underlying logic used to express beliefs. AGM requires at least a Tarskian propositional calculus. If this logic is also supra-classical and/or compact, new properties are established and the main theorems of AGM are strengthened. To model AGM it is therefore important to start by formalizing this underlying logic and its various extensions. We opted for a deep embedding in HOL which required the redefinition of all the logical operators and an axiomatization of their rules. This is certainly not efficient in terms of proof but it gives a total control of our formalization and an assurance that the logic used has no hidden properties depending on the Isabelle/HOL implementation. We use the Isabelle \<^emph>\locales\ feature and we take advantage of the inheritance/extension mechanisms between locales.\ subsection \Tarskian Logic\ text \ The first locale formalizes a Tarskian logic based on the famous Tarski's consequence operator: @{term \Cn(A)\} which gives the set of all propositions (\<^bold>\closure\) that can be inferred from the set of propositions @{term \A\}, Exactly as it is classically axiomatized in the literature, three assumptions of the locale define the consequence operator. \ locale Tarskian_logic = fixes Cn::\'a set \ 'a set\ assumes monotonicity_L: \A \ B \ Cn(A) \ Cn(B)\ and inclusion_L: \A \ Cn(A)\ and transitivity_L: \Cn(Cn(A)) \ Cn(A)\ \ \ Short notation for ``@{term \\\} can be inferred from the propositions in @{term \A\}''. \ fixes infer::\'a set \ 'a \ bool\ (infix \\\ 50) defines \A \ \ \ \ \ Cn(A)\ \ \ @{term \\\} is valid (a tautology) if it can be inferred from nothing. \ fixes valid::\'a \ bool\ (\\\) defines \\ \ \ {} \ \\ \ \ @{term \A \ \\} is all that can be infered from @{term \A\} and @{term \\\}. \ fixes expansion::\'a set \ 'a \ 'a set\ (infix \\\ 57) defines \A \ \ \ Cn(A \ {\})\ begin lemma idempotency_L: \Cn(Cn(A)) = Cn(A)\ by (simp add: inclusion_L transitivity_L subset_antisym) lemma assumption_L: \\ \ A \ A \ \\ using inclusion_L infer_def by blast lemma validD_L: \\ \ \ \ \ Cn(A)\ using monotonicity_L valid_def infer_def by fastforce lemma valid_expansion: \K = Cn(A) \ \ \ \ K \ \ = K\ by (simp add: idempotency_L insert_absorb validD_L valid_def expansion_def) lemma transitivity2_L: assumes \\\ \ B. A \ \\ and \B \ \\ shows \A \ \\ proof - from assms(1) have \B \ Cn(A)\ by (simp add: infer_def subsetI) hence \Cn(B) \ Cn(A)\ using idempotency_L monotonicity_L by blast moreover from assms(2) have \\ \ Cn(B)\ by (simp add: infer_def) ultimately show ?thesis using infer_def by blast qed lemma Cn_same: \(Cn(A) = Cn(B)) \ (\C. A \ Cn(C) \ B \ Cn(C))\ proof { assume h:\Cn(A) = Cn(B)\ from h have \\\ \ B. A \ \\ by (simp add: Tarskian_logic.assumption_L Tarskian_logic_axioms infer_def) moreover from h[symmetric] have \\\ \ A. B \ \\ by (simp add: Tarskian_logic.assumption_L Tarskian_logic_axioms infer_def) ultimately have \\C. A \ Cn(C) \ B \ Cn(C)\ using h idempotency_L inclusion_L monotonicity_L by blast } thus \Cn(A) = Cn(B) \ \C. (A \ Cn(C)) = (B \ Cn(C))\ . next { assume h:\\C. (A \ Cn(C)) = (B \ Cn(C))\ from h have \(A \ Cn(A)) = (B \ Cn(A))\ and \(A \ Cn(B)) = (B \ Cn(B))\ by simp+ hence \B \ Cn(A)\ and \A \ Cn(B)\ by (simp add: inclusion_L)+ hence \Cn(A) = Cn(B)\ using idempotency_L monotonicity_L by blast } thus \(\C. (A \ Cn(C)) = (B \ Cn(C))) \ Cn(A) = Cn(B)\ . qed \ \ The closure of the union of two consequence closures. \ lemma Cn_union: \Cn(Cn(A) \ Cn(B)) = Cn(A \ B)\ proof have \Cn(Cn(A) \ Cn(B)) \ Cn(Cn (A \ B))\ by (simp add: monotonicity_L) thus \Cn(Cn(A) \ Cn(B)) \ Cn(A \ B)\ by (simp add: idempotency_L) next have \(A \ B) \ (Cn(A) \ Cn(B))\ using inclusion_L by blast thus \Cn(A \ B) \ Cn(Cn(A) \ Cn(B))\ by (simp add: monotonicity_L) qed \ \ The closure of an infinite union of consequence closures. \ lemma Cn_Union: \Cn(\{Cn(B)|B. P B}) = Cn(\{B. P B})\ (is \?A = ?B\) proof have \?A \ Cn ?B\ apply(rule monotonicity_L, rule Union_least, auto) by (metis Sup_upper in_mono mem_Collect_eq monotonicity_L) then show \?A \ ?B\ by (simp add: idempotency_L) next show \?B \ ?A\ by (metis (mono_tags, lifting) Union_subsetI inclusion_L mem_Collect_eq monotonicity_L) qed \ \ The intersection of two closures is closed. \ lemma Cn_inter: \K = Cn(A) \ Cn(B) \ K = Cn(K)\ proof - { fix K assume h:\K = Cn(A) \ Cn(B)\ from h have \K \ Cn(A)\ and \K \ Cn(B)\ by simp+ hence \Cn(K) \ Cn(A)\ and \Cn(K) \ Cn(B)\ using idempotency_L monotonicity_L by blast+ hence \Cn(K) \ Cn(A) \ Cn(B)\ by simp with h have \K = Cn(K)\ by (simp add: inclusion_L subset_antisym) } thus \K = Cn(A) \ Cn(B) \ K = Cn(K)\ . qed \ \ An infinite intersection of closures is closed. \ lemma Cn_Inter: \K = \{Cn(B)|B. P B} \ K = Cn(K)\ proof - { fix K assume h:\K = \{Cn(B)|B. P B}\ from h have \\B. P B \ K \ Cn(B)\ by blast hence \\B. P B \ Cn(K) \ Cn(B)\ using idempotency_L monotonicity_L by blast hence \Cn(K) \ \{Cn(B)|B. P B}\ by blast with h have \K = Cn(K)\ by (simp add: inclusion_L subset_antisym) } thus \K = \{Cn(B)|B. P B} \ K = Cn(K)\ . qed end subsection \Supraclassical Logic\ text \ A Tarskian logic has only one abstract operator catching the notion of consequence. A basic case of such a logic is a \<^bold>\Supraclassical\ logic that is a logic with all classical propositional operators (e.g. conjunction (\\\), implication(\\\), negation (\\\) \dots ) together with their classical semantics. We define a new locale. In order to distinguish the propositional operators of our supraclassical logic from those of Isabelle/HOL, we use dots (e.g. \.\.\ stands for conjunction). We axiomatize the introduction and elimination rules of each operator as it is commonly established in the classical literature. As explained before, we give priority to a complete control of our logic instead of an efficient shallow embedding in Isabelle/HOL. \ locale Supraclassical_logic = Tarskian_logic + fixes true_PL:: \'a\ (\\\) and false_PL:: \'a\ (\\\) and imp_PL:: \'a \ 'a \ 'a\ (infix \.\.\ 55) and not_PL:: \'a \ 'a\ (\.\\) and conj_PL:: \'a \ 'a \ 'a\ (infix \.\.\ 55) and disj_PL:: \'a \ 'a \ 'a\ (infix \.\.\ 55) and equiv_PL:: \'a \ 'a \ 'a\ (infix \.\.\ 55) assumes true_PL: \A \ \\ and false_PL: \{\} \ p\ and impI_PL: \A \ {p} \ q \ A \ (p .\. q)\ and mp_PL: \A \ p .\. q \ A \ p \ A \ q\ and notI_PL: \A \ p .\. \ \ A \ .\ p\ and notE_PL: \A \ .\ p \ A \ (p .\. \)\ and conjI_PL: \A \ p \ A \ q \ A \ (p .\. q)\ and conjE1_PL: \A \ p .\. q \ A \ p\ and conjE2_PL: \A \ p .\. q \ A \ q\ and disjI1_PL: \A \ p \ A \ (p .\. q)\ and disjI2_PL: \A \ q \ A \ (p .\. q)\ and disjE_PL: \A \ p .\. q \ A \ p .\. r \ A \ q.\. r \ A \ r\ and equivI_PL: \A \ p .\. q \ A \ q .\. p \ A \ (p .\. q)\ and equivE1_PL: \A \ p .\. q \ A \ p .\. q\ and equivE2_PL: \A \ p .\. q \ A \ q .\. p\ \ \non intuitionistic rules\ and absurd_PL: \A \ .\ (.\ p) \ A \ p\ and ex_mid_PL: \A \ p .\. (.\ p)\ begin text \In the following, we will first retrieve the classical logic operators semantics coming from previous introduction and elimination rules\ lemma non_consistency: assumes \A \ .\ p\ and \A \ p\ shows \A \ q\ by (metis assms(1) assms(2) false_PL mp_PL notE_PL singleton_iff transitivity2_L) \ \this direct result brings directly many remarkable properties of implication (i.e. transitivity)\ lemma imp_PL: \A \ p .\. q \ A \ {p} \ q\ apply (intro iffI impI_PL) apply(rule mp_PL[where p=p], meson UnI1 assumption_L transitivity2_L) using assumption_L by auto lemma not_PL: \A \ .\ p \ A \ {p} \ \\ using notE_PL notI_PL imp_PL by blast \ \Classical logic result\ lemma notnot_PL: \A \ .\ (.\ p) \ A \ p\ apply(rule iffI, simp add:absurd_PL) by (meson mp_PL notE_PL Un_upper1 Un_upper2 assumption_L infer_def monotonicity_L not_PL singletonI subsetD) lemma conj_PL: \A \ p .\. q \ (A \ p \ A \ q)\ using conjE1_PL conjE2_PL conjI_PL by blast lemma disj_PL: \A \ p .\. q \ A \ {.\ p} \ q\ proof assume a:\A \ p .\. q\ have b:\A \ p .\. (.\ p .\. q)\ by (intro impI_PL) (meson Un_iff assumption_L insertI1 non_consistency) have c:\A \ q .\. (.\ p .\. q)\ by (simp add: assumption_L impI_PL) from a b c have \A \ .\ p .\. q\ by (erule_tac disjE_PL) simp_all then show \A \ {.\ p} \ q\ using imp_PL by blast next assume a:\A \ {.\ p} \ q\ hence b:\A \ .\ p .\. q\ by (simp add: impI_PL) then show \A \ p .\. q\ apply(rule_tac disjE_PL[OF ex_mid_PL, of A p \p .\. q\]) by (auto simp add: assumption_L disjI2_PL disjI1_PL impI_PL imp_PL) qed lemma equiv_PL:\A \ p .\. q \ (A \ {p} \ q \ A \ {q} \ p)\ using imp_PL equivE1_PL equivE2_PL equivI_PL by blast corollary valid_imp_PL: \\ (p .\. q) = ({p} \ q)\ and valid_not_PL: \\ (.\ p) = ({p} \ \)\ and valid_conj_PL: \\ (p .\. q) = (\ p \ \ q)\ and valid_disj_PL: \\ (p .\. q) = ({.\ p} \ q)\ and valid_equiv_PL:\\ (p .\. q) = ({p} \ q \ {q} \ p)\ using imp_PL not_PL conj_PL disj_PL equiv_PL valid_def by auto text\Second, we will combine each logical operator with the consequence operator \Cn\: it is a trick to profit from set theory to get many essential lemmas without complex inferences\ declare infer_def[simp] lemma nonemptyCn: \Cn(A) \ {}\ using true_PL by auto lemma Cn_true: \Cn({\}) = Cn({})\ using Cn_same true_PL by auto lemma Cn_false: \Cn({\}) = UNIV\ using false_PL by auto lemma Cn_imp: \A \ (p .\. q) \ Cn({q}) \ Cn(A \ {p})\ and Cn_imp_bis: \A \ (p .\. q) \ Cn(A \ {q}) \ Cn(A \ {p})\ using Cn_same imp_PL idempotency_L inclusion_L infer_def subset_insertI by force+ lemma Cn_not: \A \ .\ p \ Cn(A \ {p}) = UNIV\ using Cn_false Cn_imp notE_PL not_PL by fastforce lemma Cn_conj: \A \ (p .\. q) \ Cn({p}) \ Cn({q}) \ Cn(A)\ apply(intro iffI conjI_PL, frule conjE1_PL, frule conjE2_PL) using Cn_same Un_insert_right bot.extremum idempotency_L inclusion_L by auto lemma Cn_conj_bis: \Cn({p .\. q}) = Cn({p, q})\ by (unfold Cn_same) (meson Supraclassical_logic.conj_PL Supraclassical_logic_axioms insert_subset) lemma Cn_disj: \A \ (p .\. q) \ Cn({q}) \ Cn(A \ {.\ p})\ and Cn_disj_bis: \A \ (p .\. q) \ Cn(A \ {q}) \ Cn(A \ {.\ p})\ using disj_PL Cn_same imp_PL idempotency_L inclusion_L infer_def subset_insertI by force+ lemma Cn_equiv: \A \ (p .\. q) \ Cn(A \ {p}) = Cn(A \ {q})\ by (metis Cn_imp_bis equivE1_PL equivE2_PL equivI_PL set_eq_subset) corollary valid_nonemptyCn: \Cn({}) \ {}\ and valid_Cn_imp: \\ (p .\. q) \ Cn({q}) \ Cn({p})\ and valid_Cn_not: \\ (.\ p) \ Cn({p}) = UNIV\ and valid_Cn_conj: \\ (p .\. q) \ Cn({p}) \ Cn({q}) \ Cn({})\ and valid_Cn_disj: \\ (p .\. q) \ Cn({q}) \ Cn({.\ p})\ and valid_Cn_equiv: \\ (p .\. q) \ Cn({p}) = Cn({q})\ using nonemptyCn Cn_imp Cn_not Cn_conj Cn_disj Cn_equiv valid_def by auto \ \Finally, we group additional lemmas that were essential in further proofs\ lemma consistency: \Cn({p}) \ Cn({.\ p}) = Cn({})\ proof { fix q assume \{p} \ q\ and \{.\ p} \ q\ hence "{} \ p .\. q" and "{} \ (.\ p) .\. q" using impI_PL by auto hence \{} \ q\ using ex_mid_PL by (rule_tac disjE_PL[where p=p and q=\.\ p\]) blast } then show \Cn({p}) \ Cn({.\ p}) \ Cn({})\ by (simp add: subset_iff) next show \Cn({}) \ Cn({p}) \ Cn({.\ p})\ by (simp add: monotonicity_L) qed lemma Cn_notnot: \Cn({.\ (.\ \)}) = Cn({\})\ by (metis (no_types, opaque_lifting) notnot_PL valid_Cn_equiv valid_equiv_PL) lemma conj_com: \A \ p .\. q \ A \ q .\. p\ using conj_PL by auto lemma conj_com_Cn: \Cn({p .\. q}) = Cn({q .\. p})\ by (simp add: Cn_conj_bis insert_commute) lemma disj_com: \A \ p .\. q \ A \ q .\. p\ proof - { fix p q have \A \ p .\. q \ A \ q .\. p\ apply(erule disjE_PL) using assumption_L disjI2_PL disjI1_PL impI_PL by auto } then show ?thesis by auto qed lemma disj_com_Cn: \Cn({p .\. q}) = Cn({q .\. p})\ unfolding Cn_same using disj_com by simp lemma imp_contrapos: \A \ p .\. q \ A \ .\ q .\. .\ p\ by (metis Cn_not Un_insert_left Un_insert_right imp_PL notnot_PL) lemma equiv_negation: \A \ p .\. q \ A \ .\ p .\. .\ q\ using equivE1_PL equivE2_PL equivI_PL imp_contrapos by blast lemma imp_trans: \A \ p .\.q \ A \ q .\.r \ A \ p .\.r\ using Cn_imp_bis by auto lemma imp_recovery0: \A \ p .\. (p .\. q)\ apply(subst disj_PL, subst imp_contrapos) using assumption_L impI_PL by auto lemma imp_recovery1: \A \ {p .\. q} \ p \ A \ p\ using disjE_PL[OF imp_recovery0, of A p p q] assumption_L imp_PL by auto lemma imp_recovery2: \A \ p .\. q \ A \ (q .\. p) .\. p \ A \ q\ using imp_PL imp_recovery1 imp_trans by blast lemma impI2: \A \ q \ A \ p .\. q\ by (meson assumption_L impI_PL in_mono sup_ge1 transitivity2_L) lemma conj_equiv: \A \ p \ A \ ((p .\. q) .\. q)\ by (metis Un_insert_right assumption_L conjE2_PL conjI_PL equiv_PL impI2 imp_PL insertI1 sup_bot.right_neutral) lemma conj_imp: \A \ (p .\. q) .\. r \ A \ p .\. (q .\. r)\ proof assume "A \ (p .\. q) .\. r" then have "Cn (A \ {r}) \ Cn (A \ {p, q})" by (metis (no_types) Cn_conj_bis Cn_imp_bis Cn_union Un_insert_right sup_bot.right_neutral) then show \A \ p .\. (q .\. r)\ by (metis Un_insert_right impI_PL inclusion_L infer_def insert_commute insert_subset subset_eq sup_bot.right_neutral) next assume "A \ p .\. (q .\. r)" then have "A \ {p} \ {q} \ r" using imp_PL by auto then show "A \ (p .\. q) .\. r" by (metis (full_types) Cn_conj_bis Cn_union impI_PL infer_def insert_is_Un sup_assoc) qed lemma conj_not_impE_PL: \A \ (p .\. q) .\. r \ A \ (p .\. .\ q) .\. r \ A \ p .\. r\ by (meson conj_imp disjE_PL ex_mid_PL imp_PL) lemma disj_notE_PL: \A \ q \ A \ p .\. .\ q \ A \ p\ using Cn_imp Cn_imp_bis Cn_not disjE_PL notnot_PL by blast lemma disj_not_impE_PL: \A \ (p .\. q) .\. r \ A \ (p .\. .\ q) .\. r \ A \ r\ by (metis Un_insert_right disjE_PL disj_PL disj_com ex_mid_PL insert_commute sup_bot.right_neutral) lemma imp_conj: \A \ p .\. q \ A \ r .\. s \ A \ (p .\. r) .\. (q .\. s)\ apply(intro impI_PL conjI_PL, unfold imp_PL[symmetric]) by (meson assumption_L conjE1_PL conjE2_PL imp_trans infer_def insertI1 validD_L valid_imp_PL)+ lemma conj_overlap: \A \ (p .\. q) \ A \ (p .\. ((.\ p) .\. q))\ by (meson conj_PL disjI2_PL disj_com disj_notE_PL) lemma morgan: \A \ .\ (p .\. q) \ A \ (.\ p) .\. (.\ q)\ by (meson conj_imp disj_PL disj_com imp_PL imp_contrapos notE_PL notI_PL) lemma conj_superexpansion1: \A \ .\ (p .\. q) .\. .\ p \ A \ .\ p\ using conj_PL disjI1_PL morgan by auto lemma conj_superexpansion2: \A \ (p .\. q) .\. p \ A \ p\ using conj_PL disjI1_PL by auto end subsection \Compact Logic\ text\If the logic is compact, which means that any result is based on a finite set of hypothesis\ locale Compact_logic = Tarskian_logic + assumes compactness_L: \A \ \ \ (\A'. A'\ A \ finite A' \ A'\ \)\ begin text \Very important lemma preparing the application of the Zorn's lemma. It states that in a compact logic, we can not deduce \\\ if we accumulate an infinity of hypothesis groups which locally do not deduce phi\ lemma chain_closure: \\ \ \ \ subset.chain {B. \ B \ \} C \ \ \C \ \\ proof assume a:\subset.chain {B. \ B \ \} C\ and b:\\ \ \\ and \\ C \ \\ then obtain A' where c:\A'\ \ C\ and d:\finite A'\ and e:\A' \ \\ using compactness_L by blast define f where f:\f \ \a. SOME B. B \ C \ a \ B\ have g:\finite (f ` A')\ using f d by simp have h:\(f ` A') \ C\ unfolding f by auto (metis (mono_tags, lifting) Union_iff c someI_ex subset_eq) have i:\subset.chain {B. \ B \ \} (f ` A')\ using a h using a h by (meson subsetD subset_chain_def subset_trans) have \A' \ {} \ \ (f ` A') \ {B. \ B \ \}\ using g i by (meson Union_in_chain image_is_empty subset_chain_def subset_eq) hence j:\A' \ {} \ \ \(f ` A') \ \\ by simp have \A' \ \(f ` A')\ unfolding f by auto (metis (no_types, lifting) Union_iff c someI_ex subset_iff) with j e b show False by (metis infer_def monotonicity_L subsetD valid_def) qed end end diff --git a/thys/CoCon/Traceback_Properties.thy b/thys/CoCon/Traceback_Properties.thy --- a/thys/CoCon/Traceback_Properties.thy +++ b/thys/CoCon/Traceback_Properties.thy @@ -1,682 +1,682 @@ theory Traceback_Properties imports Safety_Properties begin section \Traceback properties\ text \In this section, we prove various traceback properties, by essentially giving trace-based justifications of certain occurring situations that are relevant for access to information: % \begin{description} \item{\bf Being an author. } If a user is an author of a paper, then either the user has registered the paper in the first place or, inductively, has been appointed as coauthor by another author. % \item{\bf Being a chair. } If a user is a chair of a conference, then either that user has registered the conference which has been approved by the superuser or, inductively, that user has been appointed by an existing chair of that conference. % % \item{\bf Being a PC member. } If a user is a PC member in a conference, then the user either must have been the original chair or must have been appointed by a chair. % \item{\bf Being a reviewer. } -If a user is a paper’s reviewer, then the user must have been appointed by a chair (from +If a user is a paper's reviewer, then the user must have been appointed by a chair (from among the PC members who have not declared a conflict with the paper). % \item{\bf Having conflict. } If a user has conflict with a paper, then the user is either an author of the paper or the -conflict has been declared by that user or by a paper’s author, in such a way that between +conflict has been declared by that user or by a paper's author, in such a way that between the moment when the conflict has been last declared and the current moment there is no transition that successfully removes the conflict. % \item{\bf Conference reaching a phase. } If a conference is in a given phase different from ``no phase'', then this has happened as a consequence of either a conference approval action by the superuser (if the phase is Setup) or a phase change action by a chair (otherwise). \end{description} More details and explanations can be found in \cite[Section 3.6]{cocon-JAR2021}. \ subsection \Preliminaries\ inductive trace_between :: "state \ (state,act,out) trans trace \ state \ bool" where empty[simp]: "trace_between s [] s" | step: "\trace_between s tr sh; step sh a = (ou,s')\ \ trace_between s (tr@[Trans sh a ou s']) s'" inductive_simps trace_ft_empty[simp]: "trace_between s [] s'" and trace_ft_snoc: "trace_between s (tr@[trn]) s'" thm trace_ft_empty trace_ft_snoc lemma trace_ft_append: "trace_between s (tr1@tr2) s' \ (\sh. trace_between s tr1 sh \ trace_between sh tr2 s')" apply (induction tr2 arbitrary: s' rule: rev_induct) apply simp apply (subst append_assoc[symmetric], subst trace_ft_snoc) apply (auto simp: trace_ft_snoc) done lemma trace_ft_Cons: "trace_between s (trn#tr) s' \ (\sh ou a. trn = Trans s a ou sh \ step s a = (ou,sh) \ trace_between sh tr s')" apply (subst trace_ft_append[where ?tr1.0 = "[trn]", simplified]) apply (subst trace_ft_snoc[where tr = "[]", simplified]) by auto lemmas trace_ft_simps = trace_ft_empty trace_ft_snoc trace_ft_Cons trace_ft_append inductive trace_to :: " (state,act,out) trans trace \ state \ bool" where empty: "trace_to [] istate" | step: "\trace_to tr s; step s a = (ou,s')\ \ trace_to (tr@[Trans s a ou s']) s'" lemma trace_to_ft: "trace_to tr s \ trace_between istate tr s" proof (rule,goal_cases) case 1 thus ?case by induction (auto intro: trace_between.intros) next case 2 moreover {fix s' assume "trace_between s' tr s" hence "s' = istate \ trace_to tr s" by induction (auto intro: trace_to.intros) } ultimately show ?case by auto qed inductive_simps trace_to_empty[simp]: "trace_to [] s" lemma trace_to_reach: assumes "trace_to tr s" shows "reach s" using assms apply induction apply (rule reach.intros) by (metis reach_step snd_conv) lemma reach_to_trace: assumes "reach s" obtains tr where "trace_to tr s" using assms apply (induction rule: reach_step_induct) apply (auto intro: trace_to.intros) [] by (metis surjective_pairing trace_to.step) lemma reach_trace_to_conv: "reach s \ (\tr. trace_to tr s)" by (blast intro: trace_to_reach elim: reach_to_trace) thm trace_to.induct[no_vars] lemma trace_to_induct[case_names empty step, induct set]: "\trace_to x1 x2; P [] istate; \tr s a ou s'. \trace_to tr s; P tr s; reach s; reach s'; step s a = (ou, s')\ \ P (tr ## Trans s a ou s') s'\ \ P x1 x2" apply (erule trace_to.induct) apply simp apply (frule trace_to_reach) using reach_PairI by blast subsection \Authorship\ text \ Only the creator of a paper, and users explicitly added by other authors, are authors of a paper. \ inductive isAut' :: "(state,act,out) trans trace \ confID \ userID \ paperID \ bool" where creator: "\ trn = Trans _ (Cact (cPaper cid uid _ pid _ _)) outOK _ \ \ isAut' (tr@[trn]) cid uid pid" (* "The creator of a paper is an author" *) | co_author: "\ isAut' tr cid uid' pid; trn = Trans _ (Cact (cAuthor cid uid' _ pid uid)) outOK _ \ \ isAut' (tr@[trn]) cid uid pid" (* "An author can add any other user as a coauthor" *) | irrelevant: "isAut' tr cid uid' pid \ isAut' (tr@[_]) cid uid' pid" lemma justify_author: assumes "trace_to tr s" assumes "isAut s cid uid pid" shows "isAut' tr cid uid pid" using assms proof (induction arbitrary: uid) case (empty s) thus ?case by (auto simp add: istate_def) next case (step tr s a ou s') show ?case proof (cases "isAut s cid uid pid") case True with step.IH show ?thesis by (blast intro: isAut'.intros) next case False with step.hyps step.prems obtain pass s1 s2 uid' where "a=Cact (cPaper cid uid pass pid s1 s2) \ (a=Cact (cAuthor cid uid' pass pid uid) \ isAut s cid uid' pid)" and [simp]: "ou=outOK" apply (cases a) subgoal for x1 apply (cases x1, auto simp add: c_defs) [] . subgoal for x2 apply (cases x2, auto simp add: u_defs) [] . subgoal for x3 apply (cases x3, auto simp add: uu_defs) [] . by simp_all thus ?thesis using step.IH apply (elim disjE) apply (rule isAut'.creator, auto) [] apply (rule isAut'.co_author, auto) [] done qed qed lemma author_justify: assumes "trace_to tr s" assumes "isAut' tr cid uid pid" shows "isAut s cid uid pid" using assms proof (induction arbitrary: uid) case (empty s) thus ?case by (auto simp add: istate_def elim: isAut'.cases) next case (step tr s a ou s') from step.prems show ?case proof (cases) case (creator _ _ pass s1 s2) hence [simp]: "a=Cact (cPaper cid uid pass pid s1 s2)" "ou=outOK" by simp_all from step.hyps show ?thesis by (auto simp add: c_defs) next case (co_author _ uid' _ _ pass) hence [simp]: "a=Cact (cAuthor cid uid' pass pid uid)" "ou=outOK" by simp_all from step.hyps show ?thesis by (auto simp add: c_defs) next case (irrelevant) with step.IH have AUT: "isAut s cid uid pid" by simp note roles_confIDs[OF \reach s\ AUT] with AUT \step s a = (ou, s')\ show ?thesis apply (cases a) subgoal for x1 apply (cases x1, auto simp: c_defs) [] . subgoal for x2 apply (cases x2, auto simp: u_defs) [] . subgoal for x3 apply (cases x3, auto simp: uu_defs) [] . by simp_all qed qed theorem isAut_eq: "trace_to tr s \ isAut s cid uid pid \ isAut' tr cid uid pid" (* "Trace-based equivalent of authorship" *) using justify_author author_justify by (blast) subsection \Becoming a Conference Chair\ inductive isChair' :: "(state,act,out) trans trace \ confID \ userID \ bool" where creator: "\ trn=Trans _ (Cact (cConf cid uid _ _ _)) outOK _ \ \ isChair' (tr@[trn]) cid uid" | add_chair: "\ isChair' tr cid uid'; trn = Trans _ (Cact (cChair cid uid' _ uid)) outOK _ \ \ isChair' (tr@[trn]) cid uid" | irrelevant: "\isChair' tr cid uid\ \ isChair' (tr@[_]) cid uid" lemma justify_chair: assumes "trace_to tr s" assumes "isChair s cid uid" shows "isChair' tr cid uid" using assms proof (induction arbitrary: uid) case (empty s) thus ?case by (auto simp add: istate_def) next case (step tr s a ou s') show ?case proof (cases "isChair s cid uid") case True with step.IH show ?thesis by (blast intro: isChair'.intros) next case False term cConf with step.hyps step.prems obtain pass s1 s2 uid' where "a=Cact (cConf cid uid pass s1 s2) \ (a=Cact (cChair cid uid' pass uid) \ isChair s cid uid')" and [simp]: "ou=outOK" apply (cases a) subgoal for x1 apply (cases x1, auto simp add: c_defs) [] . subgoal for x2 apply (cases x2, auto simp add: u_defs) [] . subgoal for x3 apply (cases x3, auto simp add: uu_defs) [] . by simp_all thus ?thesis using step.IH apply (elim disjE) apply (rule isChair'.creator, auto) [] apply (rule isChair'.add_chair, auto) [] done qed qed lemma chair_justify: assumes "trace_to tr s" assumes "isChair' tr cid uid" shows "isChair s cid uid" using assms proof (induction arbitrary: uid) case (empty s) thus ?case by (auto simp add: istate_def elim: isChair'.cases) next case (step tr s a ou s') from step.prems show ?case proof (cases) case (creator _ _ pass s1 s2) hence [simp]: "a=Cact (cConf cid uid pass s1 s2)" "ou=outOK" by simp_all from step.hyps show ?thesis by (auto simp add: c_defs) next case (add_chair _ uid' _ _ pass) hence [simp]: "a=Cact (cChair cid uid' pass uid)" "ou=outOK" by simp_all from step.hyps show ?thesis by (auto simp add: c_defs) next case (irrelevant) with step.IH have CH: "isChair s cid uid" by simp from CH \step s a = (ou, s')\ show ?thesis apply (cases a) subgoal for x1 apply (cases x1, auto simp: c_defs) [] . subgoal for x2 apply (cases x2, auto simp: u_defs) [] . subgoal for x3 apply (cases x3, auto simp: uu_defs) [] . by simp_all qed qed theorem isChair_eq: "trace_to tr s \ isChair s cid uid = isChair' tr cid uid" (* "Trace-based equivalent of being a chair" *) using justify_chair chair_justify by (blast) subsection \Committee Membership\ inductive isPC' :: "(state,act,out) trans trace \ confID \ userID \ bool" where chair: "isChair' tr cid uid \ isPC' tr cid uid" | add_com: "\ isChair' tr cid uid'; trn = Trans _ (Cact (cPC cid uid' _ uid)) outOK _ \ \ isPC' (tr@[trn]) cid uid" | irrelevant: "\isPC' tr cid uid\ \ isPC' (tr@[_]) cid uid" lemma justify_com: assumes "trace_to tr s" assumes "isPC s cid uid" shows "isPC' tr cid uid" using assms proof (induction arbitrary: uid) case (empty s) thus ?case by (auto simp add: istate_def) next case (step tr s a ou s') show ?case proof (cases "isPC s cid uid") case True with step.IH show ?thesis by (blast intro: isPC'.irrelevant) next case False note noPC = this show ?thesis proof (cases "isChair s' cid uid") case True thus ?thesis by (metis chair justify_chair step.hyps(1) step.hyps(4) trace_to.step) next case False note noChair=this from noPC noChair step.hyps step.prems obtain pass uid' where "(a=Cact (cPC cid uid' pass uid))" and "isChair s cid uid'" and [simp]: "ou=outOK" apply (cases a) subgoal for x1 apply (cases x1, auto simp add: c_defs) [] . subgoal for x2 apply (cases x2, auto simp add: u_defs) [] . subgoal for x3 apply (cases x3, auto simp add: uu_defs) [] . by simp_all thus ?thesis apply - apply (rule isPC'.add_com, auto simp: isChair_eq[OF \trace_to tr s\]) [] done qed qed qed lemma com_justify: assumes "trace_to tr s" assumes "isPC' tr cid uid" shows "isPC s cid uid" using assms proof (induction arbitrary: uid) case (empty s) thus ?case by (auto simp add: istate_def elim!: isPC'.cases isChair'.cases) next case (step tr s a ou s') from step.prems show ?case proof (cases) case chair thus ?thesis by (metis isChair_eq isChair_isPC step.hyps(1) step.hyps(3) step.hyps(4) trace_to.step) next case (add_com _ uid' _ _ pass) hence [simp]: "a=Cact (cPC cid uid' pass uid)" "ou=outOK" by simp_all from step.hyps show ?thesis by (auto simp add: c_defs) next case (irrelevant) with step.IH have COM: "isPC s cid uid" by simp from COM \step s a = (ou, s')\ show ?thesis apply (cases a) subgoal for x1 apply (cases x1, auto simp: c_defs) [] . subgoal for x2 apply (cases x2, auto simp: u_defs) [] . subgoal for x3 apply (cases x3, auto simp: uu_defs) [] . by simp_all qed qed theorem isPC_eq: "trace_to tr s \ isPC s cid uid = isPC' tr cid uid" (* "Trace-based equivalent of committee membership" *) using justify_com com_justify by (blast) subsection \Being a Reviewer\ inductive isRev' :: "(state,act,out) trans trace \ confID \ userID \ paperID \ bool" where add_rev: "\ isChair' tr cid uid'; trn = Trans _ (Cact (cReview cid uid' _ pid uid)) outOK _ \ \ isRev' (tr@[trn]) cid uid pid" | irrelevant: "\isRev' tr cid uid pid\ \ isRev' (tr@[_]) cid uid pid" lemma justify_rev: assumes "trace_to tr s" assumes "isRev s cid uid pid" shows "isRev' tr cid uid pid" using assms proof (induction) case empty thus ?case by (auto simp add: istate_def isRev_def) next case (step tr s a ou s') show ?case proof (cases "isRev s cid uid pid") case True with step.IH show ?thesis by (blast intro: isRev'.irrelevant) next case False note noRev = this with step.hyps step.prems obtain pass uid' where "(a=Cact (cReview cid uid' pass pid uid))" and "isChair s cid uid'" and [simp]: "ou=outOK" apply (cases a) subgoal for x1 apply (cases x1, auto simp add: c_defs isRev_def) [] . subgoal for x2 apply (cases x2, auto simp add: u_defs isRev_def) [] . subgoal for x3 apply (cases x3, auto simp add: uu_defs isRev_def) [] . by simp_all thus ?thesis apply - apply (rule isRev'.add_rev, auto simp: isChair_eq[OF \trace_to tr s\]) [] done qed qed lemma rev_justify: assumes "trace_to tr s" assumes "isRev' tr cid uid pid" shows "isRev s cid uid pid" using assms proof (induction arbitrary: uid) case (empty s) thus ?case by (auto simp add: istate_def elim!: isRev'.cases) next case (step tr s a ou s') from step.prems show ?case proof (cases) case (add_rev _ uid' _ _ pass) hence [simp]: "a=Cact (cReview cid uid' pass pid uid)" "ou=outOK" by simp_all from step.hyps show ?thesis by (auto simp add: c_defs isRev_def) next case (irrelevant) with step.IH have REV: "isRev s cid uid pid" by simp note roles_confIDs[OF step.hyps(2)] with REV \step s a = (ou, s')\ show ?thesis apply (cases a) subgoal for x1 apply (cases x1, auto simp: c_defs isRev_def) [] . subgoal for x2 apply (cases x2, auto simp: u_defs isRev_def) [] . subgoal for x3 apply (cases x3, auto simp: uu_defs isRev_def) [] . by simp_all qed qed theorem isRev_eq: "trace_to tr s \ isRev s cid uid pid = isRev' tr cid uid pid" (* "Trace-based equivalent of being a reviewer" *) using justify_rev rev_justify by (blast) subsection "Conflicts" fun irrev_conflict :: "userID \ paperID \ (state,act,out) trans \ bool" (* "Transitions causing irrevokable conflicts" *) where "irrev_conflict uid pid (Trans _ (Cact (cPaper _ uid' _ pid' _ _)) outOK _) \ uid'=uid \ pid'=pid" | "irrev_conflict uid pid (Trans _ (Cact (cAuthor _ _ _ pid' uid')) outOK _) \ uid'=uid \ pid'=pid" | "irrev_conflict uid pid _ \ False" fun set_conflict :: "userID \ paperID \ (state,act,out) trans \ bool" (* "Transitions setting conflict state, can be revoked by later reset-actions" *) where "set_conflict uid pid (Trans _ (Cact (cConflict _ _ _ pid' uid')) outOK _) \ uid'=uid \ pid'=pid" | "set_conflict uid pid (Trans _ (Uact (uPref _ uid' _ pid' Conflict)) outOK _) \ uid'=uid \ pid'=pid" | "set_conflict _ _ _ \ False" fun reset_conflict :: "userID \ paperID \ (state,act,out)trans \ bool" (* "Transitions re-setting conflict state, can be revoked by later set-actions" *) where "reset_conflict uid pid (Trans _ (Uact (uPref _ uid' _ pid' pr)) outOK _) \ uid'=uid \ pid'=pid \ pr\Conflict" | "reset_conflict _ _ _ \ False" definition conflict_trace :: "userID \ paperID \ (state,act,out) trans trace \ bool" (* "Trace that causes a conflict: It contains either an irrevokable conflict action, or the last action concerning conflicts was set-conflict" *) where "conflict_trace uid pid tr \ (\trn\set tr. irrev_conflict uid pid trn) \ (\tr1 trn tr2. tr=tr1@trn#tr2 \ set_conflict uid pid trn \ (\trn\set tr2. \reset_conflict uid pid trn))" lemma irrev_conflict_impl_author: assumes "trace_to tr s" assumes "\trn\set tr. irrev_conflict uid pid trn" shows "\cid. isAut s cid uid pid" using assms apply induction apply (auto simp add: istate_def) [] subgoal for _ _ a apply (cases a) subgoal for x1 apply (cases x1, auto simp: c_defs, (metis roles_confIDs)+) [] . subgoal for x2 apply (cases x2, auto simp: u_defs) [] . subgoal for x3 apply (cases x3, auto simp: uu_defs) [] . by simp_all done lemma irrev_conflict_impl_conflict: assumes "trace_to tr s" assumes "\trn\set tr. irrev_conflict uid pid trn" shows "pref s uid pid = Conflict" by (metis assms(1) assms(2) irrev_conflict_impl_author isAut_pref_Conflict reach_trace_to_conv) lemma conflict_justify: assumes TR: "trace_to tr s" assumes "conflict_trace uid pid tr" shows "pref s uid pid = Conflict" using assms(2) unfolding conflict_trace_def proof (cases rule: disjE[consumes 1, case_names irrev set]) case irrev thus ?thesis by (simp add: irrev_conflict_impl_conflict[OF TR]) next case set then obtain tr1 trn tr2 where [simp]: "tr = tr1 @ trn # tr2" and SET: "set_conflict uid pid trn" and NRESET: "\trn\set tr2. \ reset_conflict uid pid trn" by blast from TR obtain s1 s2 a ou where [simp]: "trn = Trans s1 a ou s2" and TR1: "trace_to tr1 s1" and STEP: "step s1 a = (ou,s2)" and TR2: "trace_between s2 tr2 s" by (fastforce simp add: trace_to_ft trace_ft_simps) from STEP SET have "pref s2 uid pid = Conflict" apply (cases a) subgoal for x1 apply (cases x1, auto simp: c_defs) [] . subgoal for x2 apply (cases x2, auto simp: u_defs) [] subgoal for _ x65 apply (cases x65, auto) [] . subgoal for _ _ _ x65 apply (cases x65, auto) [] . subgoal for _ _ _ x65 apply (cases x65, auto) [] . . by simp_all with TR2 NRESET show ?thesis apply induction subgoal by simp subgoal for _ _ _ a apply (cases a) subgoal for x1 apply (cases x1, auto simp: c_defs) [] . subgoal for x2 apply (cases x2, auto simp: u_defs) [] . subgoal for x3 apply (cases x3, auto simp: uu_defs) [] . by simp_all done qed lemma justify_conflict: assumes TR: "trace_to tr s" assumes "pref s uid pid = Conflict" shows "conflict_trace uid pid tr" using assms proof induction case empty thus ?case by (auto simp add: istate_def) next case (step tr s a ou s') let ?trn = "Trans s a ou s'" show ?case proof (cases "pref s uid pid = Conflict") case False with step.prems \step s a = (ou, s')\ have "irrev_conflict uid pid ?trn \ set_conflict uid pid ?trn" apply (cases a) subgoal for x1 apply (cases x1, auto simp: c_defs) [] . subgoal for x2 apply (cases x2, auto simp: u_defs) [] . subgoal for x3 apply (cases x3, auto simp: uu_defs) [] . by simp_all thus ?thesis unfolding conflict_trace_def by fastforce next case True with step.IH have CT: "conflict_trace uid pid tr" . from step.prems \step s a = (ou, s')\ have "\reset_conflict uid pid ?trn" apply (cases a) subgoal by simp subgoal for x2 by (cases x2, auto simp: u_defs) by simp_all thus ?thesis using CT unfolding conflict_trace_def apply clarsimp by (metis rotate1.simps(2) set_ConsD set_rotate1) qed qed theorem conflict_eq: assumes "trace_to tr s" shows "pref s uid pid = Conflict \ conflict_trace uid pid tr" using assms conflict_justify justify_conflict by auto subsection \Conference Phases\ fun is_uPhase where "is_uPhase cid (Trans _ (Uact (uConfA cid' _ _)) outOK _) \ cid'=cid" | "is_uPhase cid (Trans _ (Uact (uPhase cid' _ _ _)) outOK _) \ cid'=cid" | "is_uPhase _ _ \ False" inductive phase' :: "(state,act,out) trans trace \ confID \ nat \ bool" where initial: "phase' [] cid noPH" | approve: "\phase' tr cid noPH; trn=Trans s (Uact (uConfA cid (voronkov s) _)) outOK _ \ \ phase' (tr@[trn]) cid setPH" | advance: "\trn = (Trans _ (Uact (uPhase cid uid _ ph)) outOK _); isChair' tr cid uid\ \ phase' (tr@[trn]) cid ph" | irrelevant: "\phase' tr cid ph; \is_uPhase cid trn \ \ phase' (tr@[trn]) cid ph" lemma justify_phase: assumes "trace_to tr s" assumes "phase s cid = ph" shows "phase' tr cid ph" using assms proof (induction arbitrary: ph) case (empty s) thus ?case by (auto simp add: istate_def phase'.initial) next case (step tr s a ou s') thus ?case apply (cases a) subgoal for x1 apply (cases x1, auto simp: c_defs intro: phase'.advance phase'.irrelevant) [] . subgoal for x2 apply (cases x2, auto simp: u_defs isChair_eq intro: phase'.advance phase'.irrelevant phase'.approve, (fastforce intro: phase'.approve phase'.irrelevant phase'.advance)+ ) [] . subgoal for x3 apply (cases x3, auto simp: uu_defs intro: phase'.irrelevant) [] . by (auto intro: phase'.advance phase'.irrelevant) qed lemma phase_justify: assumes "trace_to tr s" assumes "phase' tr cid ph" shows "phase s cid = ph" using assms proof (induction arbitrary: ph) case (empty s) thus ?case by (auto simp add: istate_def elim: phase'.cases) next case (step tr s a ou s') from step.prems show ?case proof (cases) case (approve _ _ _ pass _) hence [simp]: "a=Uact (uConfA cid (voronkov s) pass)" "ou=outOK" "ph=setPH" by simp_all from step.hyps show ?thesis by (auto simp add: u_defs) next case (advance _ _ uid pass _ _) hence [simp]: "a=Uact (uPhase cid uid pass ph)" "ou=outOK" by simp_all from step.hyps show ?thesis by (auto simp add: u_defs) next case (irrelevant) with step.IH have PH: "phase s cid = ph" "\ is_uPhase cid (Trans s a ou s')" by simp_all from PH \step s a = (ou, s')\ show ?thesis apply (cases a) subgoal for x1 apply (cases x1, auto simp: c_defs) [] . subgoal for x2 apply (cases x2, auto simp: u_defs) [] . subgoal for x3 apply (cases x3, auto simp: uu_defs) [] . by simp_all qed auto qed theorem phase_eq: assumes "trace_to tr s" shows "phase s cid = ph \ phase' tr cid ph" using assms phase_justify justify_phase by blast end diff --git a/thys/CoSMeDis/Post_Confidentiality/Post_Observation_Setup_RECEIVER.thy b/thys/CoSMeDis/Post_Confidentiality/Post_Observation_Setup_RECEIVER.thy --- a/thys/CoSMeDis/Post_Confidentiality/Post_Observation_Setup_RECEIVER.thy +++ b/thys/CoSMeDis/Post_Confidentiality/Post_Observation_Setup_RECEIVER.thy @@ -1,110 +1,110 @@ theory Post_Observation_Setup_RECEIVER imports "../Safety_Properties" begin subsection \Confidentiality for a secret receiver node\ text \We verify that a group of users of a given node \j\ can learn nothing about the updates to the content of a post \PID\ located at a different node \i\ beyond the existence of an update unless \PID\ is being shared between the two nodes and one of the users is the admin at node \j\ or becomes -a remote friend of \PID\’s owner, or \PID\ is marked as public. +a remote friend of \PID\'s owner, or \PID\ is marked as public. This is formulated as a BD Security property and is proved by unwinding. See \cite{cosmedis-SandP2017} for more details. \ subsubsection\Observation setup\ (* *) type_synonym obs = "act * out" (* The observers are an arbitrary, but fixed set of users *) locale Fixed_UIDs = fixes UIDs :: "userID set" (* The secret is PID received from AID: *) locale Fixed_PID = fixes PID :: "postID" locale Fixed_AID = fixes AID :: "apiID" locale ObservationSetup_RECEIVER = Fixed_UIDs + Fixed_PID + Fixed_AID begin (* *) fun \ :: "(state,act,out) trans \ bool" where "\ (Trans _ a _ _) \ (\ uid. userOfA a = Some uid \ uid \ UIDs) \ (\ca. a = COMact ca) \ (\uid p. a = Sact (sSys uid p))" (* Note: the passwords don't really have to be purged (since identity theft is not considered in the first place); however, purging passwords looks more sane. *) (* Purging the password in starting actions: *) fun sPurge :: "sActt \ sActt" where "sPurge (sSys uid pwd) = sSys uid emptyPass" (* Purging communicating actions: user-password information is removed, and post content for PID is removed from the only kind of action that may contain such info: comReceivePost. Note: server-password info is not purged --todo: discuss this. *) fun comPurge :: "comActt \ comActt" where "comPurge (comSendServerReq uID p aID reqInfo) = comSendServerReq uID emptyPass aID reqInfo" |"comPurge (comConnectClient uID p aID sp) = comConnectClient uID emptyPass aID sp" (* *) |"comPurge (comReceivePost aID sp pID pst uID vs) = (let pst' = (if aID = AID \ pID = PID then emptyPost else pst) in comReceivePost aID sp pID pst' uID vs)" (* *) |"comPurge (comSendPost uID p aID pID) = comSendPost uID emptyPass aID pID" |"comPurge (comSendCreateOFriend uID p aID uID') = comSendCreateOFriend uID emptyPass aID uID'" |"comPurge (comSendDeleteOFriend uID p aID uID') = comSendDeleteOFriend uID emptyPass aID uID'" |"comPurge ca = ca" (* Note: No output purge here -- only for the issuer. *) fun g :: "(state,act,out)trans \ obs" where "g (Trans _ (Sact sa) ou _) = (Sact (sPurge sa), ou)" |"g (Trans _ (COMact ca) ou _) = (COMact (comPurge ca), ou)" |"g (Trans _ a ou _) = (a,ou)" lemma comPurge_simps: "comPurge ca = comSendServerReq uID p aID reqInfo \ (\p'. ca = comSendServerReq uID p' aID reqInfo \ p = emptyPass)" "comPurge ca = comReceiveClientReq aID reqInfo \ ca = comReceiveClientReq aID reqInfo" "comPurge ca = comConnectClient uID p aID sp \ (\p'. ca = comConnectClient uID p' aID sp \ p = emptyPass)" "comPurge ca = comConnectServer aID sp \ ca = comConnectServer aID sp" "comPurge ca = comReceivePost aID sp pID pst' uID v \ (\pst. ca = comReceivePost aID sp pID pst uID v \ pst' = (if pID = PID \ aID = AID then emptyPost else pst))" "comPurge ca = comSendPost uID p aID pID \ (\p'. ca = comSendPost uID p' aID pID \ p = emptyPass)" "comPurge ca = comSendCreateOFriend uID p aID uID' \ (\p'. ca = comSendCreateOFriend uID p' aID uID' \ p = emptyPass)" "comPurge ca = comReceiveCreateOFriend aID cp uID uID' \ ca = comReceiveCreateOFriend aID cp uID uID'" "comPurge ca = comSendDeleteOFriend uID p aID uID' \ (\p'. ca = comSendDeleteOFriend uID p' aID uID' \ p = emptyPass)" "comPurge ca = comReceiveDeleteOFriend aID cp uID uID' \ ca = comReceiveDeleteOFriend aID cp uID uID'" by (cases ca; auto)+ lemma g_simps: "g (Trans s a ou s') = (COMact (comSendServerReq uID p aID reqInfo), ou') \ (\p'. a = COMact (comSendServerReq uID p' aID reqInfo) \ p = emptyPass \ ou = ou')" "g (Trans s a ou s') = (COMact (comReceiveClientReq aID reqInfo), ou') \ a = COMact (comReceiveClientReq aID reqInfo) \ ou = ou'" "g (Trans s a ou s') = (COMact (comConnectClient uID p aID sp), ou') \ (\p'. a = COMact (comConnectClient uID p' aID sp) \ p = emptyPass \ ou = ou')" "g (Trans s a ou s') = (COMact (comConnectServer aID sp), ou') \ a = COMact (comConnectServer aID sp) \ ou = ou'" "g (Trans s a ou s') = (COMact (comReceivePost aID sp pID pst' uID v), ou') \ (\pst. a = COMact (comReceivePost aID sp pID pst uID v) \ pst' = (if pID = PID \ aID = AID then emptyPost else pst) \ ou = ou')" "g (Trans s a ou s') = (COMact (comSendPost uID p aID nID), O_sendPost (aid, sp, pid, pst, own, v)) \ (\p'. a = COMact (comSendPost uID p' aID nID) \ p = emptyPass \ ou = O_sendPost (aid, sp, pid, pst, own, v))" "g (Trans s a ou s') = (COMact (comSendCreateOFriend uID p aID uID'), ou') \ (\p'. a = (COMact (comSendCreateOFriend uID p' aID uID')) \ p = emptyPass \ ou = ou')" "g (Trans s a ou s') = (COMact (comReceiveCreateOFriend aID cp uID uID'), ou') \ a = COMact (comReceiveCreateOFriend aID cp uID uID') \ ou = ou'" "g (Trans s a ou s') = (COMact (comSendDeleteOFriend uID p aID uID'), ou') \ (\p'. a = COMact (comSendDeleteOFriend uID p' aID uID') \ p = emptyPass \ ou = ou')" "g (Trans s a ou s') = (COMact (comReceiveDeleteOFriend aID cp uID uID'), ou') \ a = COMact (comReceiveDeleteOFriend aID cp uID uID') \ ou = ou'" by (cases a; auto simp: comPurge_simps ObservationSetup_RECEIVER.comPurge.simps)+ end end diff --git a/thys/CoSMed/Traceback_Properties/Post_Visibility_Traceback.thy b/thys/CoSMed/Traceback_Properties/Post_Visibility_Traceback.thy --- a/thys/CoSMed/Traceback_Properties/Post_Visibility_Traceback.thy +++ b/thys/CoSMed/Traceback_Properties/Post_Visibility_Traceback.thy @@ -1,414 +1,414 @@ theory Post_Visibility_Traceback imports Traceback_Intro begin consts PID :: postID consts VIS :: vis subsection \Tracing Back Post Visibility Status\ text \We prove the following traceback property: If, at some point \t\ on a system trace, the visibility of a post \PID\ has a value \VIS\, then one of the following holds: \begin{itemize} \item Either \VIS\ is \FriendV\ (i.e., friends-only) which is the default at post creation -\item Or the post’s owner had issued a successful ``update visibility'' action setting the visibility to \VIS\, +\item Or the post's owner had issued a successful ``update visibility'' action setting the visibility to \VIS\, and no other successful update actions to \PID\'s visibility occur between the time of that action and \t\. \end{itemize} This will be captured in the predicate \proper\, and the main theorem states that \proper tr\ holds for any trace \tr\ that leads to post \PID\ acquiring visibility \VIS\. \ text \\SNC uidd trn\ means ``The transaction \trn\ is a successful post creation by user \uidd\'' \ fun SNC :: "userID \ (state,act,out) trans \ bool" where "SNC uidd (Trans s (Cact (cPost uid p pid tit)) ou s') = (ou = outOK \ (uid,pid) = (uidd,PID))" | "SNC uidd _ = False" text \\SNVU uidd vvs trn\ means "The transaction \trn\ is a successful post visibility update for \PID\, by user \uidd\, to value \vvs\'' \ fun SNVU :: "userID \ vis \ (state,act,out) trans \ bool" where "SNVU uidd vvs (Trans s (Uact (uVisPost uid p pid vs)) ou s') = (ou = outOK \ (uid,pid) = (uidd,PID) \ vs = vvs)" | "SNVU uidd vvis _ = False" definition proper :: "(state,act,out) trans trace \ bool" where "proper tr \ VIS = FriendV \ (\ uid tr1 trn tr2 trnn tr3. tr = tr1 @ trn # tr2 @ trnn # tr3 \ SNC uid trn \ SNVU uid VIS trnn \ (\ vis. never (SNVU uid vis) tr3))" (* *) definition proper1 :: "(state,act,out) trans trace \ bool" where "proper1 tr \ \ tr2 trnn tr3. tr = tr2 @ trnn # tr3 \ SNVU (owner (srcOf trnn) PID) VIS trnn" lemma not_never_ex: assumes "\ never P xs" shows "\ xs1 x xs2. xs = xs1 @ x # xs2 \ P x \ never P xs2" using assms proof(induct xs rule: rev_induct) case (Nil) thus ?case unfolding list_all_iff empty_iff by auto next case (snoc y ys) show ?case proof(cases "P y") case True thus ?thesis using snoc apply(intro exI[of _ ys]) apply(intro exI[of _ y] exI[of _ "[]"]) by auto next case False then obtain xs1 x xs2 where "ys = xs1 @ x # xs2 \ P x \ never P xs2" using snoc by auto thus ?thesis using snoc False apply(intro exI[of _ xs1]) apply(intro exI[of _ x] exI[of _ "xs2 ## y"]) by auto qed qed lemma SNVU_postIDs: assumes "validTrans trn" and "SNVU uid vs trn" shows "PID \\ postIDs (srcOf trn)" proof(cases trn) case (Trans s a ou s') then show ?thesis using assms by (cases "a") (auto simp: all_defs elim: step_elims) qed lemma SNVU_visib: assumes "validTrans trn" and "SNVU uid vs trn" shows "vis (tgtOf trn) PID = vs" proof(cases trn) case (Trans s a ou s') then show ?thesis using assms by (cases "a") (auto simp: all_defs elim: step_elims) qed lemma owner_validTrans: assumes "validTrans trn" and "PID \\ postIDs (srcOf trn)" shows "owner (srcOf trn) PID = owner (tgtOf trn) PID" proof(cases trn) case (Trans s a ou s') then show ?thesis using assms by (cases "a") (auto simp: all_defs elim: step_elims) qed lemma owner_valid: assumes "valid tr" and "PID \\ postIDs (srcOf (hd tr))" shows "owner (srcOf (hd tr)) PID = owner (tgtOf (last tr)) PID" using assms using owner_validTrans IDs_mono validTrans by induct auto lemma SNVU_vis_validTrans: assumes "validTrans trn" and "PID \\ postIDs (srcOf trn)" and "\ vs. \ SNVU (owner (srcOf trn) PID) vs trn" shows "vis (srcOf trn) PID = vis (tgtOf trn) PID" proof(cases trn) case (Trans s a ou s') then show ?thesis using assms by (cases "a") (auto simp: all_defs elim: step_elims) qed lemma SNVU_vis_valid: assumes "valid tr" and "PID \\ postIDs (srcOf (hd tr))" and "\ vis. never (SNVU (owner (srcOf (hd tr)) PID) vis) tr" shows "vis (srcOf (hd tr)) PID = vis (tgtOf (last tr)) PID" using assms proof induct case (Singl) thus ?case using SNVU_vis_validTrans by auto next case (Cons trn tr) have n: "PID \\ postIDs (srcOf (hd tr))" using Cons by (simp add: IDs_mono(2) validTrans) have v: "\ vis. never (SNVU (owner (srcOf (hd tr)) PID) vis) tr" using Cons by (simp add: owner_validTrans) have "vis (srcOf trn) PID = vis (srcOf (hd tr)) PID" using Cons SNVU_vis_validTrans by auto also have "... = vis (tgtOf (last tr)) PID" using n v Cons(4) by auto finally show ?case using Cons by auto qed lemma proper1_never: assumes vtr: "valid tr" and PID: "PID \\ postIDs (srcOf (hd tr))" and tr: "proper1 tr" and v: "vis (tgtOf (last tr)) PID = VIS" shows "\ tr2 trnn tr3. tr = tr2 @ trnn # tr3 \ SNVU (owner (srcOf trnn) PID) VIS trnn \ (\ vis. never (SNVU (owner (srcOf trnn) PID) vis) tr3)" proof- obtain tr2 trnn tr3 where tr: "tr = tr2 @ trnn # tr3" and SNVU: "SNVU (owner (srcOf trnn) PID) VIS trnn" using tr unfolding proper1_def by auto define uid where "uid \ owner (srcOf trnn) PID" show ?thesis proof(cases "never (\ trn. \ vis. SNVU uid vis trn) tr3") case True thus ?thesis using tr SNVU unfolding uid_def list_all_iff by blast next case False from not_never_ex[OF this] obtain tr3a tr3n tr3b vs where tr3: "tr3 = tr3a @ tr3n # tr3b" and SNVUtr3n: "SNVU uid vs tr3n" and n: "\ vs. never (SNVU uid vs) tr3b" unfolding list_all_iff by blast have trnn: "validTrans trnn" and tr3n: "validTrans tr3n" and vtr3: "valid tr3" using tr unfolding tr tr3 by (metis Nil_is_append_conv append_self_conv2 list.distinct(1) tr tr3 valid_ConsE valid_append vtr)+ hence PID_trnn: "PID \\ postIDs (srcOf trnn)" and PID_tr3n: "PID \\ postIDs (srcOf tr3n)" using SNVU_postIDs SNVU SNVUtr3n by auto have vvv: "valid (trnn # tr3a @ [tr3n])" using vtr unfolding tr tr3 by (smt Nil_is_append_conv append_self_conv2 hd_append2 list.distinct(1) list.sel(1) valid_Cons_iff valid_append) hence PID_tr3n': "PID \\ postIDs (tgtOf tr3n)" using tr3n SNVUtr3n by (simp add: IDs_mono(2) PID_tr3n validTrans) from owner_valid[OF vvv] PID_trnn have 000: "owner (tgtOf tr3n) PID = uid" unfolding uid_def by simp hence 0: "owner (srcOf tr3n) PID = uid" using PID_tr3n owner_validTrans tr3n by blast have 00: "vs = vis (tgtOf tr3n) PID" using SNVUtr3n tr3n SNVU_visib by auto have vis: "vs = VIS" proof(cases "tr3b = []") case True thus ?thesis using v 00 unfolding tr tr3 by simp next case False hence tgt: "tgtOf tr3n = srcOf (hd tr3b)" and tr3b: "valid tr3b" using vtr3 unfolding tr3 apply (metis valid_append list.distinct(2) self_append_conv2 valid_ConsE) by (metis False append_self_conv2 list.distinct(1) tr3 valid_Cons_iff valid_append vtr3) show ?thesis unfolding 00 tgt using v False PID_tr3n' using SNVU_vis_valid[OF tr3b _ n[unfolded 000[symmetric] tgt]] unfolding tr tr3 tgt by simp qed show ?thesis apply(intro exI[of _ "tr2 @ trnn # tr3a"]) apply(intro exI[of _ tr3n] exI[of _ tr3b]) using SNVUtr3n n unfolding tr tr3 0 vis by simp qed qed (* *) lemma SNVU_validTrans: assumes "validTrans trn" and "PID \\ postIDs (srcOf trn)" and "vis (srcOf trn) PID \ VIS" and "vis (tgtOf trn) PID = VIS" shows "SNVU (owner (srcOf trn) PID) VIS trn" proof(cases trn) case (Trans s a ou s') then show ?thesis using assms by (cases "a") (auto simp: all_defs elim: step_elims) qed lemma valid_mono_postID: assumes "valid tr" and "PID \\ postIDs (srcOf (hd tr))" shows "PID \\ postIDs (tgtOf (last tr))" using assms proof induct case (Singl trn) then show ?case using IDs_mono(2) by (cases trn) auto next case (Cons trn tr) then show ?case using IDs_mono(2) by (cases trn) auto qed lemma proper1_valid: assumes V: "VIS \ FriendV" and a: "valid tr" "PID \\ postIDs (srcOf (hd tr))" "vis (srcOf (hd tr)) PID \ VIS" "vis (tgtOf (last tr)) PID = VIS" shows "proper1 tr" using a unfolding valid_valid2 proof induct case (Singl trn) then show ?case unfolding proper1_def using SNVU_validTrans by (intro exI[of _ "owner (srcOf trn) PID"] exI[of _ "[]"] exI[of _ trn]) auto next case (Rcons tr trn) hence "PID \\ postIDs (srcOf (hd tr))" using Rcons by simp from valid_mono_postID[OF \valid2 tr\[unfolded valid2_valid] this] have "PID \\ postIDs (tgtOf (last tr))" by simp hence 0: "PID \\ postIDs (srcOf trn)" using Rcons by simp show ?case proof(cases "vis (srcOf trn) PID = VIS") case False hence "SNVU (owner (srcOf trn) PID) VIS trn" apply (intro SNVU_validTrans) using 0 Rcons by auto thus ?thesis unfolding proper1_def by (intro exI[of _ tr] exI[of _ trn] exI[of _ "[]"]) auto next case True hence "proper1 tr" using Rcons by auto then obtain trr trnn tr3 where tr: "tr = trr @ trnn # tr3" and SNVU: "SNVU (owner (srcOf trnn) PID) VIS trnn" unfolding proper1_def using V by auto have "vis (tgtOf trn) PID = VIS" using Rcons.prems by auto thus ?thesis using SNVU V unfolding proper1_def tr by(intro exI[of _ trr] exI[of _ trnn] exI[of _ "tr3 ## trn"]) auto qed qed lemma istate_postIDs: "\ PID \\ postIDs istate" unfolding istate_def by simp (* *) definition proper2 :: "(state,act,out) trans trace \ bool" where "proper2 tr \ \ uid tr1 trn tr2. tr = tr1 @ trn # tr2 \ SNC uid trn" lemma SNC_validTrans: assumes "VIS \ FriendV" and "validTrans trn" and "\ PID \\ postIDs (srcOf trn)" and "PID \\ postIDs (tgtOf trn)" shows "\ uid. SNC uid trn" proof(cases trn) case (Trans s a ou s') then show ?thesis using assms by (cases "a") (auto simp: all_defs elim: step_elims) qed lemma proper2_valid: assumes V: "VIS \ FriendV" and a: "valid tr" "\ PID \\ postIDs (srcOf (hd tr))" "PID \\ postIDs (tgtOf (last tr))" shows "proper2 tr" using a unfolding valid_valid2 proof induct case (Singl trn) then obtain uid where "SNC uid trn" using SNC_validTrans V by auto thus ?case unfolding proper2_def using SNC_validTrans by (intro exI[of _ uid] exI[of _ "[]"] exI[of _ trn]) auto next case (Rcons tr trn) show ?case proof(cases "PID \\ postIDs (srcOf trn)") case False then obtain uid where "SNC uid trn" using Rcons SNC_validTrans V by auto thus ?thesis unfolding proper2_def apply - apply (intro exI[of _ uid] exI[of _ tr]) by (intro exI[of _ trn] exI[of _ "[]"]) auto next case True hence "proper2 tr" using Rcons by auto then obtain uid tr1 trnn tr2 where tr: "tr = tr1 @ trnn # tr2" and SFRC: "SNC uid trnn" unfolding proper2_def by auto have "PID \\ postIDs (tgtOf trn)" using V Rcons.prems by auto show ?thesis using SFRC unfolding proper2_def tr apply - apply (intro exI[of _ uid] exI[of _ tr1]) by (intro exI[of _ trnn] exI[of _ "tr2 ## trn"]) simp qed qed lemma proper2_valid_istate: assumes V: "VIS \ FriendV" and a: "valid tr" "srcOf (hd tr) = istate" "PID \\ postIDs (tgtOf (last tr))" shows "proper2 tr" using proper2_valid assms istate_postIDs by auto (* *) lemma SNC_visPost: assumes "VIS \ FriendV" and "validTrans trn" "SNC uid trn" and "reach (srcOf trn)" shows "vis (tgtOf trn) PID \ VIS" proof(cases trn) case (Trans s a ou s') then show ?thesis using assms apply (cases "a") apply (auto simp: all_defs elim: step_elims) subgoal for x2 apply(cases x2) using reach_not_postIDs_vis_FriendV by (auto simp: all_defs elim: step_elims) . qed lemma SNC_postIDs: assumes "validTrans trn" and "SNC uid trn" shows "PID \\ postIDs (tgtOf trn)" proof(cases trn) case (Trans s a ou s') then show ?thesis using assms by (cases "a") (auto simp: all_defs elim: step_elims) qed lemma SNC_owner: assumes "validTrans trn" and "SNC uid trn" shows "uid = owner (tgtOf trn) PID" proof(cases trn) case (Trans s a ou s') then show ?thesis using assms by (cases "a") (auto simp: all_defs elim: step_elims) qed theorem post_accountability: assumes v: "valid tr" and i: "srcOf (hd tr) = istate" and PIDin: "PID \\ postIDs (tgtOf (last tr))" and PID: "vis (tgtOf (last tr)) PID = VIS" shows "proper tr" proof(cases "VIS = FriendV") case True thus ?thesis unfolding proper_def by auto next case False have "proper2 tr" using proper2_valid_istate[OF False v i PIDin] . then obtain uid tr1 trn trr where tr: "tr = tr1 @ trn # trr" and SNC: "SNC uid trn" unfolding proper2_def by auto hence trn: "validTrans trn" and r: "reach (srcOf trn)" using v unfolding tr apply (metis list.distinct(2) self_append_conv2 valid_ConsE valid_append) by (metis (mono_tags, lifting) append_Cons hd_append i list.sel(1) reach.simps tr v valid_append valid_init_reach) hence N: "PID \\ postIDs (tgtOf trn)" "vis (tgtOf trn) PID \ VIS" using SNC_postIDs SNC_visPost False SNC by auto hence trrNE: "trr \ []" and 1: "last tr = last trr" using PID unfolding tr by auto hence trr_v: "valid trr" using v unfolding tr by (metis valid_Cons_iff append_self_conv2 list.distinct(1) valid_append) have 0: "tgtOf trn = srcOf (hd trr)" using v trrNE unfolding tr by (metis valid_append list.distinct(2) self_append_conv2 valid_ConsE) have "proper1 trr" using proper1_valid[OF False trr_v N[unfolded 0] PID[unfolded 1]] . from proper1_never[OF trr_v N(1)[unfolded 0] this PID[unfolded 1]] obtain tr2 trnn tr3 where trr: "trr = tr2 @ trnn # tr3" and SNVU: "SNVU (owner (srcOf trnn) PID) VIS trnn" and vis: "\ vis. never (SNVU (owner (srcOf trnn) PID) vis) tr3" by auto have 00: "srcOf (hd (tr2 @ [trnn])) = tgtOf trn" using v unfolding tr trr by (metis "0" append_self_conv2 hd_append2 list.sel(1) trr) have trnn: "validTrans trnn" using trr_v unfolding trr by (metis valid_Cons_iff append_self_conv2 list.distinct(1) valid_append) have vv: "valid (tr2 @ [trnn])" using v unfolding tr trr by (smt Nil_is_append_conv append_self_conv2 hd_append2 list.distinct(1) list.sel(1) valid_Cons_iff valid_append) have "uid = owner (tgtOf trn) PID" using SNC trn SNC_owner by auto also have "... = owner (tgtOf trnn) PID" using owner_valid[OF vv] N(1) unfolding 00 by simp also have "... = owner (srcOf trnn) PID" using SNVU trnn SNVU_postIDs owner_validTrans by auto finally have uid: "uid = owner (srcOf trnn) PID" . show ?thesis unfolding proper_def apply(rule disjI2) apply(intro exI[of _ uid] exI[of _ tr1]) apply(rule exI[of _ trn], rule exI[of _ tr2]) apply(intro exI[of _ trnn] exI[of _ tr3]) using SNC SNVU vis unfolding tr trr uid by auto qed end diff --git a/thys/Complete_Non_Orders/Fixed_Points.thy b/thys/Complete_Non_Orders/Fixed_Points.thy --- a/thys/Complete_Non_Orders/Fixed_Points.thy +++ b/thys/Complete_Non_Orders/Fixed_Points.thy @@ -1,1236 +1,1236 @@ (* Author: Jérémy Dubut (2019-2021) Author: Akihisa Yamada (2019-2021) License: LGPL (see file COPYING.LESSER) *) theory Fixed_Points imports Complete_Relations begin section \Existence of Fixed Points in Complete Related Sets\ text \\label{sec:qfp-exists}\ text \The following proof is simplified and generalized from Stouti--Maaden \cite{SM13}. We construct some set whose extreme bounds -- if they exist, typically when the underlying related set is complete -- are fixed points of a monotone or inflationary function on any related set. When the related set is attractive, those are actually the least fixed points. This generalizes \cite{SM13}, relaxing reflexivity and antisymmetry.\ locale fixed_point_proof = related_set + fixes f assumes f: "f ` A \ A" begin sublocale less_eq_notations. definition AA where "AA \ {X. X \ A \ f ` X \ X \ (\Y s. Y \ X \ extreme_bound A (\) Y s \ s \ X)}" lemma AA_I: "X \ A \ f ` X \ X \ (\Y s. Y \ X \ extreme_bound A (\) Y s \ s \ X) \ X \ AA" by (unfold AA_def, safe) lemma AA_E: "X \ AA \ (X \ A \ f ` X \ X \ (\Y s. Y \ X \ extreme_bound A (\) Y s \ s \ X) \ thesis) \ thesis" by (auto simp: AA_def) definition C where "C \ \ AA" lemma A_AA: "A \ AA" by (auto intro!:AA_I f) lemma C_AA: "C \ AA" proof (intro AA_I) show "C \ A" using C_def A_AA f by auto show "f ` C \ C" unfolding C_def AA_def by auto fix B b assume B: "B \ C" "extreme_bound A (\) B b" { fix X assume X: "X \ AA" with B have "B \ X" by (auto simp: C_def) with X B have "b\X" by (auto elim!: AA_E) } then show "b \ C" by (auto simp: C_def AA_def) qed lemma CA: "C \ A" using A_AA by (auto simp: C_def) lemma fC: "f ` C \ C" using C_AA by (auto elim!: AA_E) context fixes c assumes Cc: "extreme_bound A (\) C c" begin private lemma cA: "c \ A" using Cc by auto private lemma cC: "c \ C" using Cc C_AA by (blast elim!:AA_E) private lemma fcC: "f c \ C" using cC AA_def C_AA by auto private lemma fcA: "f c \ A" using fcC CA by auto lemma qfp_as_extreme_bound: assumes infl_mono: "\x \ A. x \ f x \ (\y \ A. y \ x \ f y \ f x)" shows "f c \ c" proof (intro conjI bexI sympartpI) show "f c \ c" using fcC Cc by auto from infl_mono[rule_format, OF cA] show "c \ f c" proof (safe) text \Monotone case:\ assume mono: "\b\A. b \ c \ f b \ f c" define D where "D \ {x \ C. x \ f c}" have "D \ AA" proof (intro AA_I) show "D \ A" unfolding D_def C_def using A_AA f by auto have fxC: "x \ C \ x \ f c \ f x \ C" for x using C_AA by (auto simp: AA_def) show "f ` D \ D" proof (unfold D_def, safe intro!: fxC) fix x assume xC: "x \ C" have "x \ c" "x \ A" using Cc xC CA by auto then show "f x \ f c" using mono by (auto dest:monotoneD) qed have DC: "D \ C" unfolding D_def by auto fix B b assume BD: "B \ D" and Bb: "extreme_bound A (\) B b" have "B \ C" using DC BD by auto then have bC: "b \ C" using C_AA Bb BD by (auto elim!: AA_E) have bfc: "\a\B. a \ f c" using BD unfolding D_def by auto with f cA Bb have "b \ f c" by (auto simp: extreme_def image_subset_iff) with bC show "b \ D" unfolding D_def by auto qed then have "C \ D" unfolding C_def by auto then show "c \ f c" using cC unfolding D_def by auto qed qed lemma extreme_qfp: assumes attract: "\q \ A. \x \ A. f q \ q \ x \ f q \ x \ q" and mono: "monotone_on A (\) (\) f" shows "extreme {q \ A. f q \ q \ f q = q} (\) c" proof- have fcc: "f c \ c" apply (rule qfp_as_extreme_bound) using mono by (auto elim!: monotone_onE) define L where [simp]: "L \ {a \ A. \s \ A. (f s \ s \ f s = s) \ a \ s}" have "L \ AA" proof (unfold AA_def, intro CollectI conjI allI impI) show XA: "L \ A" by auto show "f ` L \ L" proof safe fix x assume xL: "x \ L" show "f x \ L" proof (unfold L_def, safe) have xA: "x \ A" using xL by auto then show fxA: "f x \ A" using f by auto { fix s assume sA: "s \ A" and sf: "f s \ s \ f s = s" then have "x \ s" using xL sA sf by auto then have "f x \ f s" using mono fxA sA xA by (auto elim!:monotone_onE)} note fxfs = this { fix s assume sA: "s \ A" and sf: "f s \ s" then show "f x \ s" using fxfs attract mono sf fxA sA xA by (auto elim!:monotone_onE) } { fix s assume sA: "s \ A" and sf: "f s = s" with fxfs[OF sA] show "f x \ s" by simp} qed qed fix B b assume BL: "B \ L" and b: "extreme_bound A (\) B b" then have BA: "B \ A" by auto with BL b have bA: "b \ A" by auto show "b \ L" proof (unfold L_def, safe intro!: bA) { fix s assume sA: "s \ A" and sf: "f s \ s \ f s = s" have "bound B (\) s" using sA BL b sf by auto } note Bs = this { fix s assume sA: "s \ A" and sf: "f s \ s" with b sA Bs show "b \ s" by auto } { fix s assume sA: "s \ A" and sf: "f s = s" with b sA Bs show "b \ s" by auto } qed qed then have "C \ L" by (simp add: C_def Inf_lower) with cC have "c \ L" by auto with L_def fcc show ?thesis by auto qed end lemma ex_qfp: assumes comp: "CC-complete A (\)" and C: "C \ CC" and infl_mono: "\a \ A. a \ f a \ (\b \ A. b \ a \ f b \ f a)" shows "\s \ A. f s \ s" using qfp_as_extreme_bound[OF _ infl_mono] completeD[OF comp CA C] by auto lemma ex_extreme_qfp_fp: assumes comp: "CC-complete A (\)" and C: "C \ CC" and attract: "\q \ A. \x \ A. f q \ q \ x \ f q \ x \ q" and mono: "monotone_on A (\) (\) f" shows "\c. extreme {q \ A. f q \ q \ f q = q} (\) c" using extreme_qfp[OF _ attract mono] completeD[OF comp CA C] by auto lemma ex_extreme_qfp: assumes comp: "CC-complete A (\)" and C: "C \ CC" and attract: "\q \ A. \x \ A. f q \ q \ x \ f q \ x \ q" and mono: "monotone_on A (\) (\) f" shows "\c. extreme {q \ A. f q \ q} (\) c" proof- from completeD[OF comp CA C] obtain c where Cc: "extreme_bound A (\) C c" by auto from extreme_qfp[OF Cc attract mono] have Qc: "bound {q \ A. f q \ q} (\) c" by auto have fcc: "f c \ c" apply (rule qfp_as_extreme_bound[OF Cc]) using mono by (auto simp: monotone_onD) from Cc CA have cA: "c \ A" by auto from Qc fcc cA show ?thesis by (auto intro!: exI[of _ c]) qed end context fixes less_eq :: "'a \ 'a \ bool" (infix "\" 50) and A :: "'a set" and f assumes f: "f ` A \ A" begin interpretation less_eq_notations. interpretation fixed_point_proof A "(\)" f using f by unfold_locales theorem complete_infl_mono_imp_ex_qfp: assumes comp: "UNIV-complete A (\)" and infl_mono: "\a\A. a \ f a \ (\b\A. b \ a \ f b \ f a)" shows "\s\A. f s \ s" apply (rule ex_qfp[OF comp _ infl_mono]) by auto end corollary (in antisymmetric) complete_infl_mono_imp_ex_fp: assumes comp: "UNIV-complete A (\)" and f: "f ` A \ A" and infl_mono: "\a\A. a \ f a \ (\b\A. b \ a \ f b \ f a)" shows "\s \ A. f s = s" proof- interpret less_eq_notations. from complete_infl_mono_imp_ex_qfp[OF f comp infl_mono] obtain s where sA: "s \ A" and fss: "f s \ s" by auto from f sA have fsA: "f s \ A" by auto have "f s = s" using antisym fsA sA fss by auto with sA show ?thesis by auto qed context semiattractive begin interpretation less_eq_notations. theorem complete_mono_imp_ex_extreme_qfp: assumes comp: "UNIV-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" shows "\s. extreme {p \ A. f p \ p} (\) s" proof- interpret dual: fixed_point_proof A "(\)" rewrites "dual.sym = (\)" using f by unfold_locales (auto intro!:ext) show ?thesis apply (rule dual.ex_extreme_qfp[OF complete_dual[OF comp] _ _ monotone_on_dual[OF mono]]) apply simp using f sym_order_trans by blast qed end corollary (in antisymmetric) complete_mono_imp_ex_extreme_fp: assumes comp: "UNIV-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" shows "\s. extreme {s \ A. f s = s} (\)\<^sup>- s" proof- interpret less_eq_notations. interpret fixed_point_proof A "(\)" f using f by unfold_locales have "\c. extreme {q \ A. f q \ q \ f q = q} (\) c" apply (rule ex_extreme_qfp_fp[OF comp _ _ mono]) using antisym f by (auto dest: order_sym_trans) then obtain c where c: "extreme {q \ A. f q \ q \ f q = q} (\) c" by auto then have "f c = c" using antisym f by blast with c have "extreme {q \ A. f q = q} (\) c" by auto then show ?thesis by auto qed section \Fixed Points in Well-Complete Antisymmetric Sets\ text \\label{sec:well-complete}\ text \In this section, we prove that an inflationary or monotone map over a well-complete antisymmetric set has a fixed point. In order to formalize such a theorem in Isabelle, we followed Grall's~\cite{grall10} elementary proof for Bourbaki--Witt and Markowsky's theorems. His idea is to consider well-founded derivation trees over $A$, where from a set $C \subseteq A$ of premises one can derive $f\:(\bigsqcup C)$ if $C$ is a chain. The main observation is as follows: Let $D$ be the set of all the derivable elements; that is, for each $d \in D$ there exists a well-founded derivation whose root is $d$. It is shown that $D$ is a chain, and hence one can build a derivation yielding $f\:(\bigsqcup D)$, and $f\:(\bigsqcup D)$ is shown to be a fixed point.\ lemma bound_monotone_on: assumes mono: "monotone_on A r s f" and XA: "X \ A" and aA: "a \ A" and rXa: "bound X r a" shows "bound (f`X) s (f a)" proof (safe) fix x assume xX: "x \ X" from rXa xX have "r x a" by auto with xX XA mono aA show "s (f x) (f a)" by (auto elim!:monotone_onE) qed context fixed_point_proof begin text \To avoid the usage of the axiom of choice, we carefully define derivations so that any derivable element determines its lower set. This led to the following definition:\ definition "derivation X \ X \ A \ well_ordered_set X (\) \ (\x \ X. let Y = {y \ X. y \ x} in (\y. extreme Y (\) y \ x = f y) \ f ` Y \ Y \ extreme_bound A (\) Y x)" lemma assumes "derivation P" shows derivation_A: "P \ A" and derivation_well_ordered: "well_ordered_set P (\)" using assms by (auto simp: derivation_def) lemma derivation_cases[consumes 2, case_names suc lim]: assumes "derivation X" and "x \ X" and "\Y y. Y = {y \ X. y \ x} \ extreme Y (\) y \ x = f y \ thesis" and "\Y. Y = {y \ X. y \ x} \ f ` Y \ Y \ extreme_bound A (\) Y x \ thesis" shows thesis using assms unfolding derivation_def Let_def by auto definition "derivable x \ \X. derivation X \ x \ X" lemma derivableI[intro?]: "derivation X \ x \ X \ derivable x" by (auto simp: derivable_def) lemma derivableE: "derivable x \ (\P. derivation P \ x \ P \ thesis) \ thesis" by (auto simp: derivable_def) lemma derivable_A: "derivable x \ x \ A" by (auto elim: derivableE dest:derivation_A) lemma UN_derivations_eq_derivable: "(\{P. derivation P}) = {x. derivable x}" by (auto simp: derivable_def) context assumes ord: "antisymmetric A (\)" begin interpretation antisymmetric using ord. lemma derivation_lim: assumes P: "derivation P" and fP: "f ` P \ P" and Pp: "extreme_bound A (\) P p" shows "derivation (P \ {p})" proof (cases "p \ P") case True with P show ?thesis by (auto simp: insert_absorb) next case pP: False interpret P: well_ordered_set P "(\)" using derivation_well_ordered[OF P]. have PA: "P \ A" using derivation_A[OF P]. from Pp have pA: "p \ A" by auto have bp: "bound P (\) p" using Pp by auto then have pp: "p \ p" using Pp by auto have 1: "y \ P \ {x. (x = p \ x \ P) \ x \ y} = {x \ P. x \ y}" for y using Pp by (auto dest:extreme_bound_bound) { fix x assume xP: "x \ P" and px: "p \ x" from xP Pp have "x \ p" by auto with px have "p = x" using xP PA pA by (auto intro!: antisym) with xP pP have "False" by auto } note 2 = this then have 3: "{x. (x = p \ x \ P) \ x \ p} = P" using Pp by (auto intro!: asympartpI) have wr: "well_ordered_set (P \ {p}) (\)" apply (rule well_order_extend[OF P.well_ordered_set_axioms]) using pp bp pP 2 by auto from P fP Pp show "derivation (P \ {p})" by (auto simp: derivation_def pA wr[simplified] 1 3) qed context assumes derivation_infl: "\X x y. derivation X \ x \ X \ y \ X \ x \ y \ x \ f y" and derivation_f_refl: "\X x. derivation X \ x \ X \ f x \ f x" begin lemma derivation_suc: assumes P: "derivation P" and Pp: "extreme P (\) p" shows "derivation (P \ {f p})" proof (cases "f p \ P") case True with P show ?thesis by (auto simp: insert_absorb) next case fpP: False interpret P: well_ordered_set P "(\)" using derivation_well_ordered[OF P]. have PA: "P \ A" using derivation_A[OF P]. with Pp have pP: "p \ P" and pA: "p \ A" by auto with f have fpA: "f p \ A" by auto from pP have pp: "p \ p" by auto from derivation_infl[rule_format, OF P pP pP pp] have "p \ f p". { fix x assume xP: "x \ P" then have xA: "x \ A" using PA by auto have xp: "x \ p" using xP Pp by auto from derivation_infl[rule_format, OF P xP pP this] have "x \ f p". } note Pfp = this then have bfp: "bound P (\) (f p)" by auto { fix y assume yP: "y \ P" note yfp = Pfp[OF yP] { assume fpy: "f p \ y" with yfp have "f p = y" using yP PA pA fpA antisym by auto with yP fpP have "False" by auto } with Pfp yP have "y \ f p" by auto } note Pfp = this have 1: "\y. y \ P \ {x. (x = f p \ x \ P) \ x \ y} = {x \ P. x \ y}" and 2: "{x. (x = f p \ x \ P) \ x \ f p} = P" using Pfp by auto have wr: "well_ordered_set (P \ {f p}) (\)" apply (rule well_order_extend[OF P.well_ordered_set_axioms singleton_well_ordered]) using Pfp derivation_f_refl[rule_format, OF P pP] by auto from P Pp show "derivation (P \ {f p})" by (auto simp: derivation_def wr[simplified] 1 2 fpA) qed lemma derivable_closed: assumes x: "derivable x" shows "derivable (f x)" proof (insert x, elim derivableE) fix P assume P: "derivation P" and xP: "x \ P" note PA = derivation_A[OF P] then have xA: "x \ A" using xP by auto interpret P: well_ordered_set P "(\)" using derivation_well_ordered[OF P]. interpret P.asympartp: transitive P "(\)" using P.asympartp_transitive. define Px where "Px \ {y. y \ P \ y \ x} \ {x}" then have PxP: "Px \ P" using xP by auto have "x \ x" using xP by auto then have Pxx: "extreme Px (\) x" using xP PA by (auto simp: Px_def) have wr: "well_ordered_set Px (\)" using P.well_ordered_subset[OF PxP]. { fix z y assume zPx: "z \ Px" and yP: "y \ P" and yz: "y \ z" then have zP: "z \ P" using PxP by auto have "y \ x" proof (cases "z = x") case True then show ?thesis using yz by auto next case False then have zx: "z \ x" using zPx by (auto simp: Px_def) from P.asympartp.trans[OF yz zx yP zP xP] show ?thesis. qed } then have 1: "\z. z \ Px \ {y \ Px. y \ z} = {y \ P. y \ z}" using Px_def by blast have Px: "derivation Px" using PxP PA P by (auto simp: wr derivation_def 1) from derivation_suc[OF Px Pxx] show ?thesis by (auto intro!: derivableI) qed -text \The following lemma is derived from Grall’s proof. We simplify the claim so that we +text \The following lemma is derived from Grall's proof. We simplify the claim so that we consider two elements from one derivation, instead of two derivations.\ lemma derivation_useful: assumes X: "derivation X" and xX: "x \ X" and yX: "y \ X" and xy: "x \ y" shows "f x \ y" proof- interpret X: well_ordered_set X "(\)" using derivation_well_ordered[OF X]. note XA = derivation_A[OF X] { fix x y assume xX: "x \ X" and yX: "y \ X" from xX yX have "(x \ y \ f x \ y \ f x \ X) \ (y \ x \ f y \ x \ f y \ X)" proof (induct x arbitrary: y) case (less x) note xX = \x \ X\ and IHx = this(2) with XA have xA: "x \ A" by auto from \y \ X\ show ?case proof (induct y) case (less y) note yX = \y \ X\ and IHy = this(2) with XA have yA: "y \ A" by auto show ?case proof (rule conjI; intro impI) assume xy: "x \ y" from X yX show "f x \ y \ f x \ X" proof (cases rule:derivation_cases) case (suc Z z) with XA have zX: "z \ X" and zA: "z \ A" and zy: "z \ y" and yfz: "y = f z" by auto from xX zX show ?thesis proof (cases rule: X.comparable_three_cases) case xz: less with IHy[OF zX zy] have fxz: "f x \ z" and fxX: "f x \ X" by auto from derivation_infl[rule_format, OF X fxX zX fxz] have "f x \ y" by (auto simp: yfz) with fxX show ?thesis by auto next case eq with xX zX have "x = z" by auto with yX yfz show ?thesis by auto next case zx: greater with IHy[OF zX zy] yfz xy have False by auto then show ?thesis by auto qed next case (lim Z) note Z = \Z = {z \ X. z \ y}\ and fZ = \f ` Z \ Z\ from xX xy have "x \ Z" by (auto simp: Z) with fZ have "f x \ Z" by auto then have "f x \ y" and "f x \ X" by (auto simp: Z) then show ?thesis by auto qed next assume yx: "y \ x" from X xX show "f y \ x \ f y \ X" proof (cases rule:derivation_cases) case (suc Z z) with XA have zX: "z \ X" and zA: "z \ A" and zx: "z \ x" and xfz: "x = f z" by auto from yX zX show ?thesis proof (cases rule: X.comparable_three_cases) case yz: less with IHx[OF zX zx yX] have fyz: "f y \ z" and fyX: "f y \ X" by auto from derivation_infl[rule_format, OF X fyX zX fyz] have "f y \ x" by (auto simp: xfz) with fyX show ?thesis by auto next case eq with yX zX have "y = z" by auto with xX xfz show ?thesis by auto next case greater with IHx[OF zX zx yX] xfz yx have False by auto then show ?thesis by auto qed next case (lim Z) note Z = \Z = {z \ X. z \ x}\ and fZ = \f ` Z \ Z\ from yX yx have "y \ Z" by (auto simp: Z) with fZ have "f y \ Z" by auto then have "f y \ x" and "f y \ X" by (auto simp: Z) then show ?thesis by auto qed qed qed qed } with assms show "f x \ y" by auto qed text \Next one is the main lemma of this section, stating that elements from two possibly different derivations are comparable, and moreover the lower one is in the derivation of the upper one. -The latter claim, not found in Grall’s proof, is crucial in proving that the union of all +The latter claim, not found in Grall's proof, is crucial in proving that the union of all derivations is well-related.\ lemma derivations_cross_compare: assumes X: "derivation X" and Y: "derivation Y" and xX: "x \ X" and yY: "y \ Y" shows "(x \ y \ x \ Y) \ x = y \ (y \ x \ y \ X)" proof- { fix X Y x y assume X: "derivation X" and Y: "derivation Y" and xX: "x \ X" and yY: "y \ Y" interpret X: well_ordered_set X "(\)" using derivation_well_ordered[OF X]. interpret X.asympartp: transitive X "(\)" using X.asympartp_transitive. interpret Y: well_ordered_set Y "(\)" using derivation_well_ordered[OF Y]. have XA: "X \ A" using derivation_A[OF X]. then have xA: "x \ A" using xX by auto with f have fxA: "f x \ A" by auto have YA: "Y \ A" using derivation_A[OF Y]. then have yA: "y \ A" using yY by auto with f have fyA: "f y \ A" by auto { fix Z assume Z: "Z = {z \ X. z \ x}" and fZ: "f ` Z \ Z" and Zx: "extreme_bound A (\) Z x" and IHx: "\z \ X. z \ x \ (z \ y \ z \ Y) \ z = y \ (y \ z \ y \ X)" have "(y \ x \ y \ X) \ x \ y" proof (cases "\z \ Z. y \ z") case True then obtain z where zZ: "z \ Z" and yz: "y \ z" by auto from zZ Z have zX: "z \ X" and zx: "z \ x" by auto from IHx[rule_format, OF zX zx] yz have yX: "y \ X" by auto from X.asympartp.trans[OF yz zx yX zX xX] have "y \ x". with yX show ?thesis by auto next case False have "bound Z (\) y" proof fix z assume "z \ Z" then have zX: "z \ X" and zx: "z \ x" and nyz: "\ y \ z" using Z False by auto with IHx[rule_format, OF zX zx] X show "z \ y" by auto qed with yA Zx have xy: "x \ y" by auto then show ?thesis by auto qed } note lim_any = this { fix z Z assume Z: "Z = {z \ X. z \ x}" and Zz: "extreme Z (\) z" and xfz: "x = f z" and IHx: "(z \ y \ z \ Y) \ z = y \ (y \ z \ y \ X)" have zX: "z \ X" and zx: "z \ x" using Zz Z by (auto simp: extreme_def) then have zA: "z \ A" using XA by auto from IHx have "(y \ x \ y \ X) \ x \ y" proof (elim disjE conjE) assume zy: "z \ y" and zY: "z \ Y" from derivation_useful[OF Y zY yY zy] xfz have xy: "x \ y" by auto then show ?thesis by auto next assume zy: "z = y" then have "y \ x" using zx by auto with zy zX show ?thesis by auto next assume yz: "y \ z" and yX: "y \ X" from X.asympartp.trans[OF yz zx yX zX xX] have "y \ x". with yX show ?thesis by auto qed } note lim_any this } note lim_any = this(1) and suc_any = this(2) interpret X: well_ordered_set X "(\)" using derivation_well_ordered[OF X]. interpret Y: well_ordered_set Y "(\)" using derivation_well_ordered[OF Y]. have XA: "X \ A" using derivation_A[OF X]. have YA: "Y \ A" using derivation_A[OF Y]. from xX yY show ?thesis proof (induct x arbitrary: y) case (less x) note xX = \x \ X\ and IHx = this(2) from xX XA f have xA: "x \ A" and fxA: "f x \ A" by auto from \y \ Y\ show ?case proof (induct y) case (less y) note yY = \y \ Y\ and IHy = less(2) from yY YA f have yA: "y \ A" and fyA: "f y \ A" by auto from X xX show ?case proof (cases rule: derivation_cases) case (suc Z z) note Z = \Z = {z \ X. z \ x}\ and Zz = \extreme Z (\) z\ and xfz = \x = f z\ then have zx: "z \ x" and zX: "z \ X" by auto note IHz = IHx[OF zX zx yY] have 1: "y \ x \ y \ X \ x \ y" using suc_any[OF X Y xX yY Z Zz xfz IHz] IHy by auto from Y yY show ?thesis proof (cases rule: derivation_cases) case (suc W w) note W = \W = {w \ Y. w \ y}\ and Ww = \extreme W (\) w\ and yfw = \y = f w\ then have wY: "w \ Y" and wy: "w \ y" by auto have IHw: "w \ x \ w \ X \ w = x \ x \ w \ x \ Y" using IHy[OF wY wy] by auto have "x \ y \ x \ Y \ y \ x" using suc_any[OF Y X yY xX W Ww yfw IHw] by auto with 1 show ?thesis using antisym xA yA by auto next case (lim W) note W = \W = {w \ Y. w \ y}\ and fW = \f ` W \ W\ and Wy = \extreme_bound A (\) W y\ have "x \ y \ x \ Y \ y \ x" using lim_any[OF Y X yY xX W fW Wy] IHy by auto with 1 show ?thesis using antisym xA yA by auto qed next case (lim Z) note Z = \Z = {z \ X. z \ x}\ and fZ = \f ` Z \ Z\ and Zx = \extreme_bound A (\) Z x\ have 1: "y \ x \ y \ X \ x \ y" using lim_any[OF X Y xX yY Z fZ Zx] IHx[OF _ _ yY] by auto from Y yY show ?thesis proof (cases rule: derivation_cases) case (suc W w) note W = \W = {w \ Y. w \ y}\ and Ww = \extreme W (\) w\ and yfw = \y = f w\ then have wY: "w \ Y" and wy: "w \ y" by auto have IHw: "w \ x \ w \ X \ w = x \ x \ w \ x \ Y" using IHy[OF wY wy] by auto have "x \ y \ x \ Y \ y \ x" using suc_any[OF Y X yY xX W Ww yfw IHw] by auto with 1 show ?thesis using antisym xA yA by auto next case (lim W) note W = \W = {w \ Y. w \ y}\ and fW = \f ` W \ W\ and Wy = \extreme_bound A (\) W y\ have "x \ y \ x \ Y \ y \ x" using lim_any[OF Y X yY xX W fW Wy] IHy by auto with 1 show ?thesis using antisym xA yA by auto qed qed qed qed qed interpretation derivable: well_ordered_set "{x. derivable x}" "(\)" proof (rule well_ordered_set.intro) show "antisymmetric {x. derivable x} (\)" apply unfold_locales by (auto dest: derivable_A antisym) show "well_related_set {x. derivable x} (\)" apply (fold UN_derivations_eq_derivable) apply (rule closed_UN_well_related) by (auto dest: derivation_well_ordered derivations_cross_compare well_ordered_set.axioms) qed lemmas derivable_well_ordered = derivable.well_ordered_set_axioms lemmas derivable_total_ordered = derivable.total_ordered_set_axioms lemmas derivable_well_related = derivable.well_related_set_axioms lemma pred_unique: assumes X: "derivation X" and xX: "x \ X" shows "{z \ X. z \ x} = {z. derivable z \ z \ x}" proof { fix z assume "z \ X" and "z \ x" then have "derivable z \ z \ x" using X by (auto simp: derivable_def) } then show "{z \ X. z \ x} \ {z. derivable z \ z \ x}" by auto { fix z assume "derivable z" and zx: "z \ x" then obtain Y where Y: "derivation Y" and zY: "z \ Y" by (auto simp: derivable_def) then have "z \ X" using derivations_cross_compare[OF X Y xX zY] zx by auto } then show "{z \ X. z \ x} \ {z. derivable z \ z \ x}" by auto qed text \The set of all derivable elements is itself a derivation.\ lemma derivation_derivable: "derivation {x. derivable x}" apply (unfold derivation_def) apply (safe intro!: derivable_A derivable.well_ordered_set_axioms elim!: derivableE) apply (unfold mem_Collect_eq pred_unique[symmetric]) by (auto simp: derivation_def) text \Finally, if the set of all derivable elements admits a supremum, then it is a fixed point.\ lemma assumes p: "extreme_bound A (\) {x. derivable x} p" shows sup_derivable_qfp: "f p \ p" and sup_derivable_fp: "f p = p" proof (intro antisym sympartpI) define P where "P \ {x. derivable x}" have pA: "p \ A" using p by auto have P: "derivation P" using derivation_derivable by (simp add: P_def) from derivable_closed have fP: "f ` P \ P" by (auto simp: P_def) from derivation_lim[OF P fP] p have pP: "p \ P" by (auto simp: P_def intro:derivableI) with fP have "f p \ P" by auto with p show fpp: "f p \ p" by (auto simp: P_def) show pfp: "p \ f p" apply (rule derivation_infl[rule_format, OF P]) using pP by (auto simp: P_def) from fpp pfp p f show "f p = p" by (auto intro!: antisym) qed end text "The assumptions are satisfied by monotone functions." context assumes mono: "monotone_on A (\) (\) f" begin lemma mono_imp_derivation_infl: "\X x y. derivation X \ x \ X \ y \ X \ x \ y \ x \ f y" proof (intro allI impI) fix X x y assume X: "derivation X" and xX: "x \ X" and yX: "y \ X" and xy: "x \ y" interpret X: well_ordered_set X "(\)" using derivation_well_ordered[OF X]. note XA = derivation_A[OF X] from xX yX xy show "x \ f y" proof (induct x) case (less x) note IH = this(2) and xX = \x \ X\ and yX = \y \ X\ and xy = \x \ y\ from xX yX XA have xA: "x \ A" and yA: "y \ A" by auto from X xX show ?case proof (cases rule: derivation_cases) case (suc Z z) then have zX: "z \ X" and zsx: "z \ x" and xfz: "x = f z" by auto then have zx: "z \ x" by auto from X.trans[OF zx xy zX xX yX] have zy: "z \ y". from zX XA have zA: "z \ A" by auto from zy monotone_onD[OF mono] zA yA xfz show "x \ f y" by auto next case (lim Z) note Z = \Z = {z \ X. z \ x}\ and Zx = \extreme_bound A (\) Z x\ from f yA have fyA: "f y \ A" by auto have "bound Z (\) (f y)" proof fix z assume zZ: "z \ Z" with Z xX have zsx: "z \ x" and zX: "z \ X" by auto then have zx: "z \ x" by auto from X.trans[OF zx xy zX xX yX] have zy: "z \ y". from IH[OF zX zsx yX] zy show "z \ f y" by auto qed with Zx fyA show ?thesis by auto qed qed qed lemma mono_imp_derivation_f_refl: "\X x. derivation X \ x \ X \ f x \ f x" proof (intro allI impI) fix X x assume X: "derivation X" and xX: "x \ X" interpret X: well_ordered_set X "(\)" using derivation_well_ordered[OF X]. note XA = derivation_A[OF X] from xX have "x \ x" by auto from monotone_onD[OF mono this] xX XA show "f x \ f x" by auto qed corollary mono_imp_sup_derivable_fp: assumes p: "extreme_bound A (\) {x. derivable x} p" shows "f p = p" by (simp add: sup_derivable_fp[OF mono_imp_derivation_infl mono_imp_derivation_f_refl p]) lemma mono_imp_sup_derivable_lfp: assumes p: "extreme_bound A (\) {x. derivable x} p" shows "extreme {q \ A. f q = q} (\) p" proof (safe intro!: extremeI) from p show "p \ A" by auto from sup_derivable_fp[OF mono_imp_derivation_infl mono_imp_derivation_f_refl p] show "f p = p". fix q assume qA: "q \ A" and fqq: "f q = q" have "bound {x. derivable x} (\) q" proof (safe intro!: boundI elim!:derivableE) fix x X assume X: "derivation X" and xX: "x \ X" from X interpret well_ordered_set X "(\)" by (rule derivation_well_ordered) from xX show "x \ q" proof (induct x) case (less x) note xP = this(1) and IH = this(2) with X show ?case proof (cases rule: derivation_cases) case (suc Z z) with IH[of z] have zq: "z \ q" and zX: "z \ X" by auto from monotone_onD[OF mono zq] zX qA derivation_A[OF X] show ?thesis by (auto simp: fqq suc) next case lim with IH have "bound {z \ X. z \ x} (\) q" by auto with lim qA show ?thesis by auto qed qed qed with p qA show "p \ q" by auto qed lemma mono_imp_ex_least_fp: assumes comp: "well_complete A (\)" shows "\p. extreme {q \ A. f q = q} (\) p" proof- interpret fixed_point_proof using f by unfold_locales note as = antisymmetric_axioms note infl = mono_imp_derivation_infl note refl = mono_imp_derivation_f_refl have wr: "well_related_set {x. derivable x} (\)" using derivable_well_related[OF infl refl]. have "\p. extreme_bound A (\) {x. derivable x} p" apply (rule completeD[OF comp]) using derivable_A wr by auto then obtain p where p: "extreme_bound A (\) {x. derivable x} p" by auto from p mono_imp_sup_derivable_lfp[OF p] sup_derivable_qfp[OF infl refl p] show ?thesis by auto qed end end end text \Bourbaki-Witt Theorem on well-complete pseudo-ordered set:\ theorem (in pseudo_ordered_set) well_complete_infl_imp_ex_fp: assumes comp: "well_complete A (\)" and f: "f ` A \ A" and infl: "\x \ A. \y \ A. x \ y \ x \ f y" shows "\p \ A. f p = p" proof- note as = antisymmetric_axioms interpret fixed_point_proof using f by unfold_locales have dinfl: "\X x y. derivation X \ x \ X \ y \ X \ x \ y \ x \ f y" using infl by (auto dest!:derivation_A) have drefl: "\X x. derivation X \ x \ X \ f x \ f x" using f by (auto dest!: derivation_A) have "\p. extreme_bound A (\) {x. derivable x} p" apply (rule completeD[OF comp]) using derivable_well_related[OF as dinfl drefl] derivable_A by auto with sup_derivable_fp[OF as dinfl drefl] show ?thesis by auto qed section \Completeness of (Quasi-)Fixed Points\ text \We now prove that, under attractivity, the set of quasi-fixed points is complete.\ definition setwise where "setwise r X Y \ \x\X. \y\Y. r x y" lemmas setwiseI[intro] = setwise_def[unfolded atomize_eq, THEN iffD2, rule_format] lemmas setwiseE[elim] = setwise_def[unfolded atomize_eq, THEN iffD1, elim_format, rule_format] context fixed_point_proof begin abbreviation setwise_less_eq (infix "\\<^sup>s" 50) where "(\\<^sup>s) \ setwise (\)" subsection \Least Quasi-Fixed Points for Attractive Relations.\ lemma attract_mono_imp_least_qfp: assumes attract: "attractive A (\)" and comp: "well_complete A (\)" and mono: "monotone_on A (\) (\) f" shows "\c. extreme {p \ A. f p \ p \ f p = p} (\) c \ f c \ c" proof- interpret attractive using attract by auto interpret sym: transitive A "(\)" using sym_transitive. define ecl ("[_]\<^sub>\") where "[x]\<^sub>\ \ {y \ A. x \ y} \ {x}" for x define Q where "Q \ {[x]\<^sub>\ |. x \ A}" { fix X x assume XQ: "X \ Q" and xX: "x \ X" then have XA: "X \ A" by (auto simp: Q_def ecl_def) then have xA: "x \ A" using xX by auto obtain q where qA: "q \ A" and X: "X = [q]\<^sub>\" using XQ by (auto simp: Q_def) have xqqx: "x \ q \ x = q" using X xX by (auto simp: ecl_def) {fix y assume yX: "y \ X" then have yA: "y \ A" using XA by auto have "y \ q \ y = q" using yX X by (auto simp: ecl_def) then have "x \ y \ y = x" using sym_order_trans xqqx xA qA yA by blast } then have 1: "X \ [x]\<^sub>\" using X qA by (auto simp: ecl_def) { fix y assume "y \ A" and "x \ y \ y = x" then have "q \ y \ y = q" using sym_order_trans xqqx xA qA by blast } then have 2: "X \ [x]\<^sub>\" using X xX by (auto simp: ecl_def) from 1 2 have "X = [x]\<^sub>\" by auto } then have XQx: "\X \ Q. \x \ X. X = [x]\<^sub>\" by auto have RSLE_eq: "X \ Q \ Y \ Q \ x \ X \ y \ Y \ x \ y \ X \\<^sup>s Y" for X Y x y proof- assume XQ: "X \ Q" and YQ: "Y \ Q" and xX: "x \ X" and yY: "y \ Y" and xy: "x \ y" then have XA: "X \ A" and YA: "Y \ A" by (auto simp: Q_def ecl_def) then have xA: "x \ A" and yA: "y \ A" using xX yY by auto { fix xp yp assume xpX: "xp \ X" and ypY: "yp \ Y" then have xpA: "xp \ A" and ypA: "yp \ A" using XA YA by auto then have "xp \ x \ xp = x" using xpX XQx xX XQ by (auto simp: ecl_def) then have xpy: "xp \ y" using attract[OF _ _ xy xpA xA yA] xy by blast have "yp \ y \ yp = y" using ypY XQx yY YQ by (auto simp: ecl_def) then have "xp \ yp" using dual.attract[OF _ _ xpy ypA yA xpA] xpy by blast } then show "X \\<^sup>s Y" using XQ YQ XA YA by auto qed have compQ: "well_complete Q (\\<^sup>s)" proof (intro completeI, safe) fix XX assume XXQ: "XX \ Q" and XX: "well_related_set XX (\\<^sup>s)" have BA: "\XX \ A" using XXQ by (auto simp: Q_def ecl_def) from XX interpret XX: well_related_set XX "(\\<^sup>s)". interpret UXX: semiattractive "\XX" "(\)" by (rule semiattractive_subset[OF BA]) have "well_related_set (\XX) (\)" proof(unfold_locales) fix Y assume YXX: "Y \ \XX" and Y0: "Y \ {}" have "{X \ XX. X \ Y \ {}} \ {}" using YXX Y0 by auto from XX.nonempty_imp_ex_extreme[OF _ this] obtain E where E: "extreme {X \ XX. X \ Y \ {}} (\\<^sup>s)\<^sup>- E" by auto then have "E \ Y \ {}" by auto then obtain e where eE: "e \ E" and eX: "e \ Y" by auto have "extreme Y (\) e" proof (intro extremeI eX) fix x assume xY: "x \ Y" with YXX obtain X where XXX: "X \ XX" and xX: "x \ X" by auto with xY E XXX have "E \\<^sup>s X" by auto with eE xX show "e \ x" by auto qed then show "\e. extreme Y (\) e" by auto qed with completeD[OF comp BA] obtain b where extb: "extreme_bound A (\) (\XX) b" by auto then have bb: "b \ b" using extreme_def bound_def by auto have bA: "b \ A" using extb extreme_def by auto then have XQ: "[b]\<^sub>\ \ Q" using Q_def bA by auto have bX: "b \ [b]\<^sub>\" by (auto simp: ecl_def) have "extreme_bound Q (\\<^sup>s) XX [b]\<^sub>\" proof(intro extreme_boundI) show "[b]\<^sub>\ \ Q" using XQ. next fix Y assume YXX: "Y \ XX" then have YQ: "Y \ Q" using XXQ by auto then obtain y where yA: "y \ A" and Yy: "Y = [y]\<^sub>\" by (auto simp: Q_def) then have yY: "y \ Y" by (auto simp: ecl_def) then have "y \ \XX" using yY YXX by auto then have "y \ b" using extb by auto then show "Y \\<^sup>s [b]\<^sub>\" using RSLE_eq[OF YQ XQ yY bX] by auto next fix Z assume boundZ: "bound XX (\\<^sup>s) Z" and ZQ: "Z \ Q" then obtain z where zA: "z \ A" and Zz: "Z = [z]\<^sub>\" by (auto simp: Q_def) then have zZ: "z \ Z" by (auto simp: ecl_def) { fix y assume "y \ \XX" then obtain Y where yY: "y \ Y" and YXX: "Y \ XX" by auto then have YA: "Y \ A" using XXQ Q_def by (auto simp: ecl_def) then have "Y \\<^sup>s Z" using YXX boundZ bound_def by auto then have "y \ z" using yY zZ by auto } then have "bound (\XX) (\) z" by auto then have "b \ z" using extb zA by auto then show "[b]\<^sub>\ \\<^sup>s Z" using RSLE_eq[OF XQ ZQ bX zZ] by auto qed then show "Ex (extreme_bound Q (\\<^sup>s) XX)" by auto qed interpret Q: antisymmetric Q "(\\<^sup>s)" proof fix X Y assume XY: "X \\<^sup>s Y" and YX: "Y \\<^sup>s X" and XQ: "X \ Q" and YQ: "Y \ Q" then obtain q where qA: "q \ A" and X: "X = [q]\<^sub>\" using Q_def by auto then have qX: "q \ X" using X by (auto simp: ecl_def) then obtain p where pA: "p \ A" and Y: "Y = [p]\<^sub>\" using YQ Q_def by auto then have pY: "p \ Y" using X by (auto simp: ecl_def) have pq: "p \ q" using XQ YQ YX qX pY by auto have "q \ p" using XQ YQ XY qX pY by auto then have "p \ X" using pq X pA by (auto simp: ecl_def) then have "X = [p]\<^sub>\" using XQ XQx by auto then show "X = Y" using Y by (auto simp: ecl_def) qed define F where "F X \ {y \ A. \x \ X. y \ f x} \ f ` X" for X have XQFXQ: "\X. X \ Q \ F X \ Q" proof- fix X assume XQ: "X \ Q" then obtain x where xA: "x \ A" and X: "X = [x]\<^sub>\" using Q_def by auto then have xX: "x \ X" by (auto simp: ecl_def) have fxA: "f x \ A" using xA f by auto have FXA: "F X \ A" using f fxA X by (auto simp: F_def ecl_def) have "F X = [f x]\<^sub>\" proof (unfold X, intro equalityI subsetI) fix z assume zFX: "z \ F [x]\<^sub>\" then obtain y where yX: "y \ [x]\<^sub>\" and zfy: "z \ f y \ z = f y" by (auto simp: F_def) have yA: "y \ A" using yX xA by (auto simp: ecl_def) with f have fyA: "f y \ A" by auto have zA: "z \ A" using zFX FXA by (auto simp: X) have "y \ x \ y = x" using X yX by (auto simp: ecl_def) then have "f y \ f x \ f y = f x" using mono xA yA by (auto simp: monotone_on_def) then have "z \ f x \ z = f x" using zfy sym.trans[OF _ _ zA fyA fxA] by (auto simp:) with zA show "z \ [f x]\<^sub>\" by (auto simp: ecl_def) qed (auto simp: xX F_def ecl_def) with FXA show "F X \ Q" by (auto simp: Q_def ecl_def) qed then have F: "F ` Q \ Q" by auto then interpret Q: fixed_point_proof Q "(\\<^sup>s)" F by unfold_locales have monoQ: "monotone_on Q (\\<^sup>s) (\\<^sup>s) F" proof (intro monotone_onI) fix X Y assume XQ: "X \ Q" and YQ: "Y \ Q" and XY: "X \\<^sup>s Y" then obtain x y where xX: "x \ X" and yY: "y \ Y" using Q_def by (auto simp: ecl_def) then have xA: "x \ A" and yA: "y \ A" using XQ YQ by (auto simp: Q_def ecl_def) have "x \ y" using XY xX yY by auto then have fxfy: "f x \ f y" using monotone_on_def[of A "(\)" "(\)" f] xA yA mono by auto have fxgX: "f x \ F X" using xX F_def by blast have fygY: "f y \ F Y" using yY F_def by blast show "F X \\<^sup>s F Y" using RSLE_eq[OF XQFXQ[OF XQ] XQFXQ[OF YQ] fxgX fygY fxfy]. qed have QdA: "{x. Q.derivable x} \ Q" using Q.derivable_A by auto note asQ = Q.antisymmetric_axioms note infl = Q.mono_imp_derivation_infl[OF asQ monoQ] note f_refl = Q.mono_imp_derivation_f_refl[OF asQ monoQ] from Q.mono_imp_ex_least_fp[OF asQ monoQ compQ] obtain P where P: "extreme {q \ Q. F q = q} (\\<^sup>s)\<^sup>- P" by auto then have PQ: "P \ Q" by (auto simp: extreme_def) from P have FPP: "F P = P" using PQ by auto with P have PP: "P \\<^sup>s P" by auto from P obtain p where pA: "p \ A" and Pp: "P = [p]\<^sub>\" using Q_def by auto then have pP: "p \ P" by (auto simp: ecl_def) then have fpA: "f p \ A" using pA f by auto have "f p \ F P" using pP F_def fpA by auto then have "F P = [f p]\<^sub>\" using XQx XQFXQ[OF PQ] by auto then have fp: "f p \ p \ f p = p" using pP FPP by (auto simp: ecl_def) have "p \ p" using PP pP by auto with fp have fpp: "f p \ p" by auto have e: "extreme {p \ A. f p \ p \ f p = p} (\) p" proof (intro extremeI CollectI conjI pA fp, elim CollectE conjE) fix q assume qA: "q \ A" and fq: "f q \ q \ f q = q" define Z where "Z \ {z \ A. q \ z}\{q}" then have qZ: "q \ Z" using qA by auto then have ZQ: "Z \ Q" using qA by (auto simp: Z_def Q_def ecl_def) have fqA: "f q \ A" using qA f by auto then have "f q \ Z" using fq by (auto simp: Z_def) then have 1: "Z = [f q]\<^sub>\" using XQx ZQ by auto then have "f q \ F Z" using qZ fqA by (auto simp: F_def) then have "F Z = [f q]\<^sub>\" using XQx XQFXQ[OF ZQ] by auto with 1 have "Z = F Z" by auto then have "P \\<^sup>s Z" using P ZQ by auto then show "p \ q" using pP qZ by auto qed with fpp show ?thesis using e by auto qed subsection \General Completeness\ lemma attract_mono_imp_fp_qfp_complete: assumes attract: "attractive A (\)" and comp: "CC-complete A (\)" and wr_CC: "\C \ A. well_related_set C (\) \ C \ CC" and extend: "\X \ CC. \Y \ CC. X \\<^sup>s Y \ X \ Y \ CC" and mono: "monotone_on A (\) (\) f" and P: "P \ {x \ A. f x = x}" shows "CC-complete ({q \ A. f q \ q} \ P) (\)" proof (intro completeI) interpret attractive using attract. fix X assume Xfix: "X \ {q \ A. f q \ q} \ P" and XCC: "X \ CC" with P have XA: "X \ A" by auto define B where "B \ {b \ A. \a \ X. a \ b}" { fix s a assume sA: "s \ A" and as: "\a \ X. a \ s" and aX: "a \ X" then have aA: "a \ A" using XA by auto then have fafs: "f a \ f s" using mono f aX sA as by (auto elim!: monotone_onE) have "a \ f s" proof (cases "f a = a") case True then show ?thesis using fafs by auto next case False then have "a \ f a" using P aX Xfix by auto also from fafs have "f a \ f s" by auto finally show ?thesis using f aA sA by auto qed } with f have fBB: "f ` B \ B" unfolding B_def by auto have BA: "B \ A" by (auto simp: B_def) have compB: "CC-complete B (\)" proof (unfold complete_def, intro allI impI) fix Y assume YS: "Y \ B" and YCC: "Y \ CC" with BA have YA: "Y \ A" by auto define C where "C \ X\Y" then have CA: "C \ A" using XA YA C_def by auto have XY: "X \\<^sup>s Y" using B_def YS by auto then have CCC: "C \ CC" using extend XA YA XCC YCC C_def by auto then obtain s where s: "extreme_bound A (\) C s" using completeD[OF comp CA CCC] by auto then have sA: "s \ A" by auto show "Ex (extreme_bound B (\) Y)" proof (intro exI extreme_boundI) { fix x assume "x \ X" then have "x \ s" using s C_def by auto } then show "s \ B" using sA B_def by auto next fix y assume "y \ Y" then show "y \ s" using s C_def using extremeD by auto next fix c assume cS: "c \ B" and "bound Y (\) c" then have "bound C (\) c" using C_def B_def by auto then show "s \ c" using s BA cS by auto qed qed from fBB interpret B: fixed_point_proof B "(\)" f by unfold_locales from BA have *: "{x \ A. f x \ x} \ B = {x \ B. f x \ x}" by auto have asB: "attractive B (\)" using attractive_subset[OF BA] by auto have monoB: "monotone_on B (\) (\) f" using monotone_on_cmono[OF BA] mono by (auto dest!: le_funD) have compB: "well_complete B (\)" using wr_CC compB BA by (simp add: complete_def) from B.attract_mono_imp_least_qfp[OF asB compB monoB] obtain l where "extreme {p \ B. f p \ p \ f p = p} (\) l" and fll: "f l \ l" by auto with P have l: "extreme ({p \ B. f p \ p} \ P \ B) (\) l" by auto show "Ex (extreme_bound ({q \ A. f q \ q} \ P) (\) X)" proof (intro exI extreme_boundI) show "l \ {q \ A. f q \ q} \ P" using l BA by auto fix a assume "a \ X" with l show "a \ l" by (auto simp: B_def) next fix c assume c: "bound X (\) c" and cfix: "c \ {q \ A. f q \ q} \ P" with P have cA: "c \ A" by auto with c have "c \ B" by (auto simp: B_def) with cfix l show "l \ c" by auto qed qed lemma attract_mono_imp_qfp_complete: assumes "attractive A (\)" and "CC-complete A (\)" and "\C \ A. well_related_set C (\) \ C \ CC" and "\X \ CC. \Y \ CC. X \\<^sup>s Y \ X \ Y \ CC" and "monotone_on A (\) (\) f" shows "CC-complete {p \ A. f p \ p} (\)" using attract_mono_imp_fp_qfp_complete[OF assms, of "{}"] by simp lemma antisym_mono_imp_fp_complete: assumes anti: "antisymmetric A (\)" and comp: "CC-complete A (\)" and wr_CC: "\C \ A. well_related_set C (\) \ C \ CC" and extend: "\X \ CC. \Y \ CC. X \\<^sup>s Y \ X \ Y \ CC" and mono: "monotone_on A (\) (\) f" shows "CC-complete {p \ A. f p = p} (\)" proof- interpret antisymmetric using anti. have *: "{q \ A. f q \ q} \ {p \ A. f p = p}" using f by (auto intro!: antisym) from * attract_mono_imp_fp_qfp_complete[OF attractive_axioms comp wr_CC extend mono, of "{p\A. f p = p}"] show ?thesis by (auto simp: subset_Un_eq) qed end subsection \Instances\ subsubsection \Instances under attractivity\ context attractive begin interpretation less_eq_notations. text \Full completeness\ theorem mono_imp_qfp_UNIV_complete: assumes comp: "UNIV-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" shows "UNIV-complete {p \ A. f p \ p} (\)" apply (intro fixed_point_proof.attract_mono_imp_qfp_complete comp mono) apply unfold_locales by (auto simp: f) text \Connex completeness\ theorem mono_imp_qfp_connex_complete: assumes comp: "{X. connex X (\)}-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" shows "{X. connex X (\)}-complete {p \ A. f p \ p} (\)" apply (intro fixed_point_proof.attract_mono_imp_qfp_complete mono comp) apply unfold_locales by (auto simp: f intro: connex_union well_related_set.connex_axioms) text \Directed completeness\ theorem mono_imp_qfp_directed_complete: assumes comp: "{X. directed X (\)}-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" shows "{X. directed X (\)}-complete {p \ A. f p \ p} (\)" apply (intro fixed_point_proof.attract_mono_imp_qfp_complete mono comp) apply unfold_locales by (auto simp: f intro!: directed_extend intro: well_related_set.connex_axioms connex.directed) text \Well Completeness\ theorem mono_imp_qfp_well_complete: assumes comp: "well_complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" shows "well_complete {p \ A. f p \ p} (\)" apply (intro fixed_point_proof.attract_mono_imp_qfp_complete mono comp) apply unfold_locales by (auto simp: f well_related_extend) end subsubsection \Usual instances under antisymmetry \ context antisymmetric begin text \Knaster--Tarski\ theorem mono_imp_fp_complete: assumes comp: "UNIV-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" shows "UNIV-complete {p \ A. f p = p} (\)" proof- interpret fixed_point_proof using f by unfold_locales show ?thesis apply (intro antisym_mono_imp_fp_complete mono antisymmetric_axioms comp) by auto qed text \Markowsky 1976\ theorem mono_imp_fp_connex_complete: assumes comp: "{X. connex X (\)}-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" shows "{X. connex X (\)}-complete {p \ A. f p = p} (\)" proof- interpret fixed_point_proof using f by unfold_locales show ?thesis apply (intro antisym_mono_imp_fp_complete antisymmetric_axioms mono comp) by (auto intro: connex_union well_related_set.connex_axioms) qed text \Pataraia\ theorem mono_imp_fp_directed_complete: assumes comp: "{X. directed X (\)}-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" shows "{X. directed X (\)}-complete {p \ A. f p = p} (\)" proof- interpret fixed_point_proof using f by unfold_locales show ?thesis apply (intro antisym_mono_imp_fp_complete mono antisymmetric_axioms comp) by (auto intro: directed_extend connex.directed well_related_set.connex_axioms) qed text \Bhatta \& George 2011\ theorem mono_imp_fp_well_complete: assumes comp: "well_complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" shows "well_complete {p \ A. f p = p} (\)" proof- interpret fixed_point_proof using f by unfold_locales show ?thesis apply (intro antisym_mono_imp_fp_complete mono antisymmetric_axioms comp) by (auto intro!: antisym well_related_extend) qed end end diff --git a/thys/Complex_Geometry/More_Transcendental.thy b/thys/Complex_Geometry/More_Transcendental.thy --- a/thys/Complex_Geometry/More_Transcendental.thy +++ b/thys/Complex_Geometry/More_Transcendental.thy @@ -1,224 +1,224 @@ (* ---------------------------------------------------------------------------- *) section \Introduction\ (* ---------------------------------------------------------------------------- *) text \The complex plane or some of its parts (e.g., the unit disc or the upper half plane) are often taken as the domain in which models of various geometries (both Euclidean and non-Euclidean ones) are formalized. The complex plane gives simpler and more compact formulas than the Cartesian plane. Within complex plane is easier to describe geometric objects and perform the calculations (usually shedding some new light on the subject). We give a formalization of the extended complex plane (given both as a complex projective space and as the Riemann sphere), its objects (points, circles and lines), and its transformations (Möbius transformations).\ (* ---------------------------------------------------------------------------- *) section \Related work\ (* ---------------------------------------------------------------------------- *) text\During the last decade, there have been many results in formalizing -geometry in proof-assistants. Parts of Hilbert’s seminal book +geometry in proof-assistants. Parts of Hilbert's seminal book ,,Foundations of Geometry'' \cite{hilbert} have been formalized both in Coq and Isabelle/Isar. Formalization of first two groups of axioms in Coq, in an intuitionistic setting was done by Dehlinger et al. \cite{hilbert-coq}. First formalization in Isabelle/HOL was done by Fleuriot and Meikele \cite{hilbert-isabelle}, and some further developments were made in master thesis of Scott \cite{hilbert-scott}. Large fragments of Tarski's geometry \cite{tarski} have been formalized in Coq by Narboux et al. \cite{narboux-tarski}. Within Coq, -there are also formalizations of von Plato’s constructive geometry by +there are also formalizations of von Plato's constructive geometry by Kahn \cite{vonPlato,von-plato-formalization}, French high school geometry by Guilhot \cite{guilhot} and ruler and compass geometry by Duprat \cite{duprat2008}, etc. In our previous work \cite{petrovic2012formalizing}, we have already formally investigated a Cartesian model of Euclidean geometry. \ (* ---------------------------------------------------------------------------- *) section \Background theories\ (* ---------------------------------------------------------------------------- *) text \In this section we introduce some basic mathematical notions and prove some lemmas needed in the rest of our formalization. We describe: \<^item> trigonometric functions, \<^item> complex numbers, \<^item> systems of two and three linear equations with two unknowns (over arbitrary fields), \<^item> quadratic equations (over real and complex numbers), systems of quadratic and real equations, and systems of two quadratic equations, \<^item> two-dimensional vectors and matrices over complex numbers. \ (* -------------------------------------------------------------------------- *) subsection \Library Additions for Trigonometric Functions\ (* -------------------------------------------------------------------------- *) theory More_Transcendental imports Complex_Main "HOL-Library.Periodic_Fun" begin text \Additional properties of @{term sin} and @{term cos} functions that are later used in proving conjectures for argument of complex number.\ text \Sign of trigonometric functions on some characteristic intervals.\ lemma cos_lt_zero_on_pi2_pi [simp]: assumes "x > pi/2" and "x \ pi" shows "cos x < 0" using cos_gt_zero_pi[of "pi - x"] assms by simp text \Value of trigonometric functions in points $k\pi$ and $\frac{\pi}{2} + k\pi$.\ lemma sin_kpi [simp]: fixes k::int shows "sin (k * pi) = 0" by (simp add: sin_zero_iff_int2) lemma cos_odd_kpi [simp]: fixes k::int assumes "odd k" shows "cos (k * pi) = -1" by (simp add: assms mult.commute) lemma cos_even_kpi [simp]: fixes k::int assumes "even k" shows "cos (k * pi) = 1" by (simp add: assms mult.commute) lemma sin_pi2_plus_odd_kpi [simp]: fixes k::int assumes "odd k" shows "sin (pi / 2 + k * pi) = -1" using assms by (simp add: sin_add) lemma sin_pi2_plus_even_kpi [simp]: fixes k::int assumes "even k" shows "sin (pi / 2 + k * pi) = 1" using assms by (simp add: sin_add) text \Solving trigonometric equations and systems with special values (0, 1, or -1) of sine and cosine functions\ lemma cos_0_iff_canon: assumes "cos \ = 0" and "-pi < \" and "\ \ pi" shows "\ = pi/2 \ \ = -pi/2" by (smt (verit, best) arccos_0 arccos_cos assms cos_minus divide_minus_left) lemma sin_0_iff_canon: assumes "sin \ = 0" and "-pi < \" and "\ \ pi" shows "\ = 0 \ \ = pi" using assms sin_eq_0_pi by force lemma cos0_sin1: assumes "sin \ = 1" shows "\ k::int. \ = pi/2 + 2*k*pi" by (smt (verit, ccfv_threshold) assms cos_diff cos_one_2pi_int cos_pi_half mult_cancel_right1 sin_pi_half sin_plus_pi) (* TODO: add lemmas for cos = -1, sin = 0 and cos = 0, sin = -1 *) text \Sine is injective on $[-\frac{\pi}{2}, \frac{\pi}{2}]$\ lemma sin_inj: assumes "-pi/2 \ \ \ \ \ pi/2" and "-pi/2 \ \' \ \' \ pi/2" assumes "\ \ \'" shows "sin \ \ sin \'" by (metis assms divide_minus_left sin_inj_pi) text \Periodicity of trigonometric functions\ text \The following are available in HOL-Decision\_Procs.Approximation\_Bounds, but we want to avoid that dependency\ lemma sin_periodic_nat [simp]: fixes n :: nat shows "sin (x + n * (2 * pi)) = sin x" by (metis (no_types, opaque_lifting) add.commute add.left_neutral cos_2npi cos_one_2pi_int mult.assoc mult.commute mult.left_neutral mult_zero_left sin_add sin_int_2pin) lemma sin_periodic_int [simp]: fixes i :: int shows "sin (x + i * (2 * pi)) = sin x" by (metis add.right_neutral cos_int_2pin mult.commute mult.right_neutral mult_zero_right sin_add sin_int_2pin) lemma cos_periodic_nat [simp]: fixes n :: nat shows "cos (x + n * (2 * pi)) = cos x" by (metis add.left_neutral cos_2npi cos_add cos_periodic mult.assoc mult_2 mult_2_right of_nat_numeral sin_periodic sin_periodic_nat) lemma cos_periodic_int [simp]: fixes i :: int shows "cos (x + i * (2 * pi)) = cos x" by (metis cos_add cos_int_2pin diff_zero mult.commute mult.right_neutral mult_zero_right sin_int_2pin) text \Values of both sine and cosine are repeated only after multiples of $2\cdot \pi$\ lemma sin_cos_eq: fixes a b :: real assumes "cos a = cos b" and "sin a = sin b" shows "\ k::int. a - b = 2*k*pi" by (metis assms cos_diff cos_one_2pi_int mult.commute sin_cos_squared_add3) text \The following two lemmas are consequences of surjectivity of cosine for the range $[-1, 1]$.\ lemma ex_cos_eq: assumes "-pi/2 \ \ \ \ \ pi/2" assumes "a \ 0" and "a < 1" shows "\ \'. -pi/2 \ \' \ \' \ pi/2 \ \' \ \ \ cos (\ - \') = a" proof- have "arccos a > 0" "arccos a \ pi/2" using \a \ 0\ \a < 1\ using arccos_lt_bounded arccos_le_pi2 by auto show ?thesis proof (cases "\ - arccos a \ - pi/2") case True thus ?thesis using assms \arccos a > 0\ \arccos a \ pi/2\ by (rule_tac x = "\ - arccos a" in exI) auto next case False thus ?thesis using assms \arccos a > 0\ \arccos a \ pi/2\ by (rule_tac x = "\ + arccos a" in exI) auto qed qed lemma ex_cos_gt: assumes "-pi/2 \ \ \ \ \ pi/2" assumes "a < 1" shows "\ \'. -pi/2 \ \' \ \' \ pi/2 \ \' \ \ \ cos (\ - \') > a" proof- obtain a' where "a' \ 0" "a' > a" "a' < 1" by (metis assms(2) dense_le_bounded linear not_one_le_zero) thus ?thesis using ex_cos_eq[of \ a'] assms by auto qed text \The function @{term atan2} is a generalization of @{term arctan} that takes a pair of coordinates of non-zero points returns its angle in the range $[-\pi, \pi)$.\ definition atan2 where "atan2 y x = (if x > 0 then arctan (y/x) else if x < 0 then if y > 0 then arctan (y/x) + pi else arctan (y/x) - pi else if y > 0 then pi/2 else if y < 0 then -pi/2 else 0)" lemma atan2_bounded: shows "-pi \ atan2 y x \ atan2 y x < pi" using arctan_bounded[of "y/x"] zero_le_arctan_iff[of "y/x"] arctan_le_zero_iff[of "y/x"] zero_less_arctan_iff[of "y/x"] arctan_less_zero_iff[of "y/x"] using divide_neg_neg[of y x] divide_neg_pos[of y x] divide_pos_pos[of y x] divide_pos_neg[of y x] unfolding atan2_def by (simp (no_asm_simp)) auto end diff --git a/thys/Delta_System_Lemma/Cardinal_Library.thy b/thys/Delta_System_Lemma/Cardinal_Library.thy --- a/thys/Delta_System_Lemma/Cardinal_Library.thy +++ b/thys/Delta_System_Lemma/Cardinal_Library.thy @@ -1,708 +1,708 @@ section\Cardinal Arithmetic under Choice\label{sec:cardinal-lib}\ theory Cardinal_Library imports ZF_Library ZF.Cardinal_AC begin text\This theory includes results on cardinalities that depend on $\AC$\ subsection\Results on cardinal exponentiation\ text\Non trivial instances of cardinal exponentiation require that the relevant function spaces are well-ordered, hence this implies a strong use of choice.\ lemma cexp_eqpoll_cong: assumes "A \ A'" "B \ B'" shows "A\<^bsup>\B\<^esup> = A'\<^bsup>\B'\<^esup>" unfolding cexp_def using cardinal_eqpoll_iff function_space_eqpoll_cong assms by simp lemma cexp_cexp_cmult: "(\\<^bsup>\\1\<^esup>)\<^bsup>\\2\<^esup> = \\<^bsup>\\2 \ \1\<^esup>" proof - have "(\\<^bsup>\\1\<^esup>)\<^bsup>\\2\<^esup> = (\1 \ \)\<^bsup>\\2\<^esup>" using cardinal_eqpoll by (intro cexp_eqpoll_cong) (simp_all add:cexp_def) also have " \ = \\<^bsup>\\2 \ \1\<^esup>" unfolding cexp_def using curry_eqpoll cardinal_cong by blast also have " \ = \\<^bsup>\\2 \ \1\<^esup>" using cardinal_eqpoll[THEN eqpoll_sym] unfolding cmult_def by (intro cexp_eqpoll_cong) (simp) finally show ?thesis . qed lemma cardinal_Pow: "|Pow(X)| = 2\<^bsup>\X\<^esup>" \ \Perhaps it's better with |X|\ using cardinal_eqpoll_iff[THEN iffD2, OF Pow_eqpoll_function_space] unfolding cexp_def by simp lemma cantor_cexp: assumes "Card(\)" shows "\ < 2\<^bsup>\\\<^esup>" using assms Card_is_Ord Card_cexp proof (intro not_le_iff_lt[THEN iffD1] notI) assume "2\<^bsup>\\\<^esup> \ \" then have "|Pow(\)| \ \" using cardinal_Pow by simp with assms have "Pow(\) \ \" using cardinal_eqpoll_iff Card_le_imp_lepoll Card_cardinal_eq by auto then obtain g where "g \ inj(Pow(\), \)" by blast then show "False" using cantor_inj by simp qed simp lemma cexp_left_mono: assumes "\1 \ \2" shows "\1\<^bsup>\\\<^esup> \ \2\<^bsup>\\\<^esup>" (* \ \short, unreadable proof: \ unfolding cexp_def using subset_imp_lepoll[THEN lepoll_imp_cardinal_le] assms le_subset_iff[THEN iffD1, OF assms] Pi_weaken_type[of _ _ "\_. \1" "\_. \2"] by auto *) proof - from assms have "\1 \ \2" using le_subset_iff by simp then have "\ \ \1 \ \ \ \2" using Pi_weaken_type by auto then show ?thesis unfolding cexp_def using lepoll_imp_cardinal_le subset_imp_lepoll by simp qed lemma cantor_cexp': assumes "2 \ \" "Card(\)" shows "\ < \\<^bsup>\\\<^esup>" using cexp_left_mono assms cantor_cexp lt_trans2 by blast lemma InfCard_cexp: assumes "2 \ \" "InfCard(\)" shows "InfCard(\\<^bsup>\\\<^esup>)" using assms cantor_cexp'[THEN leI] le_trans Card_cexp unfolding InfCard_def by auto lemmas InfCard_cexp' = InfCard_cexp[OF nats_le_InfCard, simplified] \ \\<^term>\InfCard(\) \ InfCard(\) \ InfCard(\\<^bsup>\\\<^esup>)\\ subsection\Miscellaneous\ lemma cardinal_RepFun_le: "|{f(a) . a\A}| \ |A|" proof - have "(\x\A. f(x)) \ surj(A, {f(a) . a\A})" unfolding surj_def using lam_funtype by auto then show ?thesis using surj_implies_cardinal_le by blast qed lemma subset_imp_le_cardinal: "A \ B \ |A| \ |B|" using subset_imp_lepoll[THEN lepoll_imp_cardinal_le] . lemma lt_cardinal_imp_not_subset: "|A| < |B| \ \ B \ A" using subset_imp_le_cardinal le_imp_not_lt by blast lemma cardinal_lt_csucc_iff: "Card(K) \ |K'| < K\<^sup>+ \ |K'| \ K" by (simp add: Card_lt_csucc_iff) lemma cardinal_UN_le_nat: "(\i. i\\ \ |X(i)| \ \) \ |\i\\. X(i)| \ \" by (simp add: cardinal_UN_le InfCard_nat) lemma leqpoll_imp_cardinal_UN_le: notes [dest] = InfCard_is_Card Card_is_Ord assumes "InfCard(K)" "J \ K" "\i. i\J \ |X(i)| \ K" shows "|\i\J. X(i)| \ K" proof - from \J \ K\ obtain f where "f \ inj(J,K)" by blast define Y where "Y(k) \ if k\range(f) then X(converse(f)`k) else 0" for k have "i\J \ f`i \ K" for i using inj_is_fun[OF \f \ inj(J,K)\] by auto have "(\i\J. X(i)) \ (\i\K. Y(i))" proof (standard, elim UN_E) fix x i assume "i\J" "x\X(i)" with \f \ inj(J,K)\ \i\J \ f`i \ K\ have "x \ Y(f`i)" "f`i \ K" unfolding Y_def using inj_is_fun[OF \f \ inj(J,K)\] right_inverse apply_rangeI by auto then show "x \ (\i\K. Y(i))" by auto qed then have "|\i\J. X(i)| \ |\i\K. Y(i)|" unfolding Y_def using subset_imp_le_cardinal by simp with assms \\i. i\J \ f`i \ K\ show "|\i\J. X(i)| \ K" using inj_converse_fun[OF \f \ inj(J,K)\] unfolding Y_def by (rule_tac le_trans[OF _ cardinal_UN_le]) (auto intro:Ord_0_le)+ qed lemma cardinal_lt_csucc_iff': includes Ord_dests assumes "Card(\)" shows "\ < |X| \ \\<^sup>+ \ |X|" using assms cardinal_lt_csucc_iff[of \ X] Card_csucc[of \] not_le_iff_lt[of "\\<^sup>+" "|X|"] not_le_iff_lt[of "|X|" \] by blast lemma lepoll_imp_subset_bij: "X \ Y \ (\Z. Z \ Y \ Z \ X)" proof assume "X \ Y" then obtain j where "j \ inj(X,Y)" by blast then have "range(j) \ Y" "j \ bij(X,range(j))" using inj_bij_range inj_is_fun range_fun_subset_codomain by blast+ then show "\Z. Z \ Y \ Z \ X" using eqpoll_sym unfolding eqpoll_def by force next assume "\Z. Z \ Y \ Z \ X" then obtain Z f where "f \ bij(Z,X)" "Z \ Y" unfolding eqpoll_def by force then have "converse(f) \ inj(X,Y)" using bij_is_inj inj_weaken_type bij_converse_bij by blast then show "X \ Y" by blast qed text\The following result proves to be very useful when combining \<^term>\cardinal\ and \<^term>\eqpoll\ in a calculation.\ lemma cardinal_Card_eqpoll_iff: "Card(\) \ |X| = \ \ X \ \" using Card_cardinal_eq[of \] cardinal_eqpoll_iff[of X \] by auto \ \Compare @{thm "le_Card_iff"}\ lemma lepoll_imp_lepoll_cardinal: assumes "X \ Y" shows "X \ |Y|" using assms cardinal_Card_eqpoll_iff[of "|Y|" Y] lepoll_eq_trans[of _ _ "|Y|"] by simp lemma lepoll_Un: assumes "InfCard(\)" "A \ \" "B \ \" shows "A \ B \ \" proof - have "A \ B \ sum(A,B)" using Un_lepoll_sum . moreover note assms moreover from this have "|sum(A,B)| \ \ \ \" using sum_lepoll_mono[of A \ B \] lepoll_imp_cardinal_le unfolding cadd_def by auto ultimately show ?thesis using InfCard_cdouble_eq Card_cardinal_eq InfCard_is_Card Card_le_imp_lepoll[of "sum(A,B)" \] lepoll_trans[of "A\B"] by auto qed lemma cardinal_Un_le: assumes "InfCard(\)" "|A| \ \" "|B| \ \" shows "|A \ B| \ \" using assms lepoll_Un le_Card_iff InfCard_is_Card by auto text\This is the unconditional version under choice of @{thm Cardinal.Finite_cardinal_iff}.\ lemma Finite_cardinal_iff': "Finite(|i|) \ Finite(i)" using cardinal_eqpoll_iff eqpoll_imp_Finite_iff by fastforce lemma cardinal_subset_of_Card: assumes "Card(\)" "a \ \" shows "|a| < \ \ |a| = \" proof - from assms have "|a| < |\| \ |a| = |\|" using subset_imp_le_cardinal le_iff by simp with assms show ?thesis using Card_cardinal_eq by simp qed lemma cardinal_cases: includes Ord_dests shows "Card(\) \ |X| < \ \ \ |X| \ \" using not_le_iff_lt by auto subsection\Countable and uncountable sets\ \ \Kunen's Definition I.10.5\ definition countable :: "i\o" where "countable(X) \ X \ \" lemma countableI[intro]: "X \ \ \ countable(X)" unfolding countable_def by simp lemma countableD[dest]: "countable(X) \ X \ \" unfolding countable_def by simp lemma countable_iff_cardinal_le_nat: "countable(X) \ |X| \ \" using le_Card_iff[of \ X] Card_nat unfolding countable_def by simp lemma lepoll_countable: "X \ Y \ countable(Y) \ countable(X)" using lepoll_trans[of X Y] by blast \ \Next lemma can be proved without using AC\ lemma surj_countable: "countable(X) \ f \ surj(X,Y) \ countable(Y)" using surj_implies_cardinal_le[of f X Y, THEN le_trans] countable_iff_cardinal_le_nat by simp lemma Finite_imp_countable: "Finite(X) \ countable(X)" unfolding Finite_def by (auto intro:InfCard_nat nats_le_InfCard[of _ \, THEN le_imp_lepoll] dest!:eq_lepoll_trans[of X _ \]) lemma countable_imp_countable_UN: assumes "countable(J)" "\i. i\J \ countable(X(i))" shows "countable(\i\J. X(i))" using assms leqpoll_imp_cardinal_UN_le[of \ J X] InfCard_nat countable_iff_cardinal_le_nat by auto lemma countable_union_countable: assumes "\x. x \ C \ countable(x)" "countable(C)" shows "countable(\C)" using assms countable_imp_countable_UN[of C "\x. x"] by simp abbreviation uncountable :: "i\o" where "uncountable(X) \ \ countable(X)" lemma uncountable_iff_nat_lt_cardinal: "uncountable(X) \ \ < |X|" using countable_iff_cardinal_le_nat not_le_iff_lt by simp lemma uncountable_not_empty: "uncountable(X) \ X \ 0" using empty_lepollI by auto lemma uncountable_imp_Infinite: "uncountable(X) \ Infinite(X)" using uncountable_iff_nat_lt_cardinal[of X] lepoll_nat_imp_Infinite[of X] cardinal_le_imp_lepoll[of \ X] leI by simp lemma uncountable_not_subset_countable: assumes "countable(X)" "uncountable(Y)" shows "\ (Y \ X)" using assms lepoll_trans subset_imp_lepoll[of Y X] by blast subsection\Results on Alephs\ lemma nat_lt_Aleph1: "\ < \\<^bsub>1\<^esub>" by (simp add: Aleph_def lt_csucc) lemma zero_lt_Aleph1: "0 < \\<^bsub>1\<^esub>" by (rule lt_trans[of _ "\"], auto simp add: ltI nat_lt_Aleph1) lemma le_aleph1_nat: "Card(k) \ k<\\<^bsub>1\<^esub> \ k \ \" by (simp add: Aleph_def Card_lt_csucc_iff Card_nat) lemma Aleph_succ: "\\<^bsub>succ(\)\<^esub> = \\<^bsub>\\<^esub>\<^sup>+" unfolding Aleph_def by simp lemma lesspoll_aleph_plus_one: assumes "Ord(\)" shows "d \ \\<^bsub>succ(\)\<^esub> \ d \ \\<^bsub>\\<^esub>" using assms lesspoll_csucc Aleph_succ Card_is_Ord by simp lemma cardinal_Aleph [simp]: "Ord(\) \ |\\<^bsub>\\<^esub>| = \\<^bsub>\\<^esub>" using Card_cardinal_eq by simp \ \Could be proved without using AC\ lemma Aleph_lesspoll_increasing: includes Aleph_intros shows "a < b \ \\<^bsub>a\<^esub> \ \\<^bsub>b\<^esub>" using cardinal_lt_iff_lesspoll[of "\\<^bsub>a\<^esub>" "\\<^bsub>b\<^esub>"] Card_cardinal_eq[of "\\<^bsub>b\<^esub>"] lt_Ord lt_Ord2 Card_Aleph[THEN Card_is_Ord] by auto lemma uncountable_iff_subset_eqpoll_Aleph1: includes Ord_dests notes Aleph_zero_eq_nat[simp] Card_nat[simp] Aleph_succ[simp] shows "uncountable(X) \ (\S. S \ X \ S \ \\<^bsub>1\<^esub>)" proof assume "uncountable(X)" then have "\\<^bsub>1\<^esub> \ X" using uncountable_iff_nat_lt_cardinal cardinal_lt_csucc_iff' cardinal_le_imp_lepoll by force then obtain S where "S \ X" "S \ \\<^bsub>1\<^esub>" using lepoll_imp_subset_bij by auto then show "\S. S \ X \ S \ \\<^bsub>1\<^esub>" using cardinal_cong Card_csucc[of \] Card_cardinal_eq by auto next assume "\S. S \ X \ S \ \\<^bsub>1\<^esub>" then have "\\<^bsub>1\<^esub> \ X" using subset_imp_lepoll[THEN [2] eq_lepoll_trans, of "\\<^bsub>1\<^esub>" _ X, OF eqpoll_sym] by auto then show "uncountable(X)" using Aleph_lesspoll_increasing[of 0 1, THEN [2] lesspoll_trans1, of "\\<^bsub>1\<^esub>"] lepoll_trans[of "\\<^bsub>1\<^esub>" X \] by auto qed lemma lt_Aleph_imp_cardinal_UN_le_nat: "function(G) \ domain(G) \ \ \ \n\domain(G). |G`n|<\\<^bsub>1\<^esub> \ |\n\domain(G). G`n|\\" proof - assume "function(G)" let ?N="domain(G)" and ?R="\n\domain(G). G`n" assume "?N \ \" assume Eq1: "\n\?N. |G`n|<\\<^bsub>1\<^esub>" { fix n assume "n\?N" with Eq1 have "|G`n| \ \" using le_aleph1_nat by simp } then have "n\?N \ |G`n| \ \" for n . with \?N \ \\ show ?thesis using InfCard_nat leqpoll_imp_cardinal_UN_le by simp qed lemma Aleph1_eq_cardinal_vimage: "f:\\<^bsub>1\<^esub>\\ \ \n\\. |f-``{n}| = \\<^bsub>1\<^esub>" proof - assume "f:\\<^bsub>1\<^esub>\\" then have "function(f)" "domain(f) = \\<^bsub>1\<^esub>" "range(f)\\" by (simp_all add: domain_of_fun fun_is_function range_fun_subset_codomain) let ?G="\n\range(f). f-``{n}" from \f:\\<^bsub>1\<^esub>\\\ have "range(f) \ \" by (simp add: range_fun_subset_codomain) then have "domain(?G) \ \" using subset_imp_lepoll by simp have "function(?G)" by (simp add:function_lam) from \f:\\<^bsub>1\<^esub>\\\ have "n\\ \ f-``{n} \ \\<^bsub>1\<^esub>" for n using Pi_vimage_subset by simp with \range(f) \ \\ have "\\<^bsub>1\<^esub> = (\n\range(f). f-``{n})" proof (intro equalityI, intro subsetI) fix x assume "x \ \\<^bsub>1\<^esub>" with \f:\\<^bsub>1\<^esub>\\\ \function(f)\ \domain(f) = \\<^bsub>1\<^esub>\ have "x \ f-``{f`x}" "f`x \ range(f)" using function_apply_Pair vimage_iff apply_rangeI by simp_all then show "x \ (\n\range(f). f-``{n})" by auto qed auto { assume "\n\range(f). |f-``{n}| < \\<^bsub>1\<^esub>" then have "\n\domain(?G). |?G`n| < \\<^bsub>1\<^esub>" using zero_lt_Aleph1 by (auto) with \function(?G)\ \domain(?G) \ \\ have "|\n\domain(?G). ?G`n|\\" using lt_Aleph_imp_cardinal_UN_le_nat by blast then have "|\n\range(f). f-``{n}|\\" by simp with \\\<^bsub>1\<^esub> = _\ have "|\\<^bsub>1\<^esub>| \ \" by simp then have "\\<^bsub>1\<^esub> \ \" using Card_Aleph Card_cardinal_eq by simp then have "False" using nat_lt_Aleph1 by (blast dest:lt_trans2) } with \range(f)\\\ obtain n where "n\\" "\(|f -`` {n}| < \\<^bsub>1\<^esub>)" by blast moreover from this have "\\<^bsub>1\<^esub> \ |f-``{n}|" using not_lt_iff_le Card_is_Ord by auto moreover note \n\\ \ f-``{n} \ \\<^bsub>1\<^esub>\ ultimately show ?thesis using subset_imp_le_cardinal[THEN le_anti_sym, of _ "\\<^bsub>1\<^esub>"] Card_Aleph Card_cardinal_eq by auto qed \ \There is some asymmetry between assumptions and conclusion (\<^term>\(\)\ versus \<^term>\cardinal\)\ lemma eqpoll_Aleph1_cardinal_vimage: assumes "X \ \\<^bsub>1\<^esub>" "f : X \ \" shows "\n\\. |f-``{n}| = \\<^bsub>1\<^esub>" proof - from assms obtain g where "g\bij(\\<^bsub>1\<^esub>,X)" using eqpoll_sym by blast with \f : X \ \\ have "f O g : \\<^bsub>1\<^esub> \ \" "converse(g) \ bij(X, \\<^bsub>1\<^esub>)" using bij_is_fun comp_fun bij_converse_bij by blast+ then obtain n where "n\\" "|(f O g)-``{n}| = \\<^bsub>1\<^esub>" using Aleph1_eq_cardinal_vimage by auto then have "\\<^bsub>1\<^esub> = |converse(g) `` (f -``{n})|" using image_comp converse_comp unfolding vimage_def by simp also from \converse(g) \ bij(X, \\<^bsub>1\<^esub>)\ \f: X\ \\ have "\ = |f -``{n}|" using range_of_subset_eqpoll[of "converse(g)" X _ "f -``{n}"] bij_is_inj cardinal_cong bij_is_fun eqpoll_sym Pi_vimage_subset by fastforce finally show ?thesis using \n\\\ by auto qed subsection\Applications of transfinite recursive constructions\ definition rec_constr :: "[i,i] \ i" where "rec_constr(f,\) \ transrec(\,\a g. f`(g``a))" text\The function \<^term>\rec_constr\ allows to perform \<^emph>\recursive constructions\: given a choice function on the powerset of some set, a transfinite sequence is created by successively choosing some new element. The next result explains its use.\ lemma rec_constr_unfold: "rec_constr(f,\) = f`({rec_constr(f,\). \\\})" using def_transrec[OF rec_constr_def, of f \] image_lam by simp lemma rec_constr_type: assumes "f:Pow(G)\ G" "Ord(\)" shows "rec_constr(f,\) \ G" using assms(2,1) by (induct rule:trans_induct) (subst rec_constr_unfold, rule apply_type[of f "Pow(G)" "\_. G"], auto) text\The next lemma is an application of recursive constructions. It works under the assumption that whenever the already constructed subsequence is small enough, another element can be added.\ lemma bounded_cardinal_selection: includes Ord_dests assumes "\X. |X| < \ \ X \ G \ \a\G. \s\X. Q(s,a)" "b\G" "Card(\)" shows "\S. S : \ \ G \ (\\ \ \. \\ \ \. \<\ \ Q(S`\,S`\))" proof - - let ?cdlt\="{X\Pow(G) . |X|<\}" \ \“cardinal less than \<^term>\\\”\ + let ?cdlt\="{X\Pow(G) . |X|<\}" \ \``cardinal less than \<^term>\\\''\ and ?inQ="\Y.{a\G. \s\Y. Q(s,a)}" from assms have "\Y \ ?cdlt\. \a. a \ ?inQ(Y)" by blast then have "\f. f \ Pi(?cdlt\,?inQ)" using AC_ball_Pi[of ?cdlt\ ?inQ] by simp then obtain f where f_type:"f \ Pi(?cdlt\,?inQ)" by auto moreover define Cb where "Cb \ \_\Pow(G)-?cdlt\. b" moreover from \b\G\ have "Cb \ Pow(G)-?cdlt\ \ G" unfolding Cb_def by simp moreover note \Card(\)\ ultimately have "f \ Cb : (\x\Pow(G). ?inQ(x) \ G)" using fun_Pi_disjoint_Un[ of f ?cdlt\ ?inQ Cb "Pow(G)-?cdlt\" "\_.G"] Diff_partition[of "{X\Pow(G). |X|<\}" "Pow(G)", OF Collect_subset] by auto moreover have "?inQ(x) \ G = G" for x by auto ultimately have "f \ Cb : Pow(G) \ G" by simp define S where "S\\\\\. rec_constr(f \ Cb, \)" from \f \ Cb: Pow(G) \ G\ \Card(\)\ have "S : \ \ G" using Ord_in_Ord unfolding S_def by (intro lam_type rec_constr_type) auto moreover have "\\\\. \\\\. \ < \ \ Q(S ` \, S ` \)" proof (intro ballI impI) fix \ \ assume "\\\" with \Card(\)\ have "{rec_constr(f \ Cb, x) . x\\} = {S`x . x \ \}" using Ord_trans[OF _ _ Card_is_Ord, of _ \ \] unfolding S_def by auto moreover from \\\\\ \S : \ \ G\ \Card(\)\ have "{S`x . x \ \} \ G" using Ord_trans[OF _ _ Card_is_Ord, of _ \ \] apply_type[of S \ "\_. G"] by auto moreover from \Card(\)\ \\\\\ have "|{S`x . x \ \}| < \" using cardinal_RepFun_le[of \] Ord_in_Ord lt_trans1[of "|{S`x . x \ \}|" "|\|" \] Card_lt_iff[THEN iffD2, of \ \, OF _ _ ltI] by force moreover have "\x\\. Q(S`x, f ` {S`x . x \ \})" proof - from calculation and f_type have "f ` {S`x . x \ \} \ {a\G. \x\\. Q(S`x,a)}" using apply_type[of f ?cdlt\ ?inQ "{S`x . x \ \}"] by blast then show ?thesis by simp qed moreover assume "\\\" "\ < \" moreover note \\\\\ \Cb \ Pow(G)-?cdlt\ \ G\ ultimately show "Q(S ` \, S ` \)" using fun_disjoint_apply1[of "{S`x . x \ \}" Cb f] domain_of_fun[of Cb] ltD[of \ \] by (subst (2) S_def, auto) (subst rec_constr_unfold, auto) qed ultimately show ?thesis by blast qed text\The following basic result can, in turn, be proved by a bounded-cardinal selection.\ lemma Infinite_iff_lepoll_nat: "Infinite(X) \ \ \ X" proof assume "Infinite(X)" then obtain b where "b\X" using Infinite_not_empty by auto { fix Y assume "|Y| < \" then have "Finite(Y)" using Finite_cardinal_iff' ltD nat_into_Finite by blast with \Infinite(X)\ have "X \ Y" by auto } with \b\X\ obtain S where "S : \ \ X" "\\\\. \\\\. \ < \ \ S`\ \ S`\" using bounded_cardinal_selection[of \ X "\x y. x\y"] Card_nat by blast moreover from this have "\ \ \ \ \ \ \ \ \\\ \ S`\ \ S`\" for \ \ by (rule_tac lt_neq_symmetry[of "\" "\\ \. S`\ \ S`\"]) auto ultimately show "\ \ X" unfolding lepoll_def inj_def by blast qed (intro lepoll_nat_imp_Infinite) lemma Infinite_InfCard_cardinal: "Infinite(X) \ InfCard(|X|)" using lepoll_eq_trans eqpoll_sym lepoll_nat_imp_Infinite Infinite_iff_lepoll_nat Inf_Card_is_InfCard cardinal_eqpoll by simp lemma Finite_to_one_surj_imp_cardinal_eq: assumes "F \ Finite_to_one(X,Y) \ surj(X,Y)" "Infinite(X)" shows "|Y| = |X|" proof - from \F \ Finite_to_one(X,Y) \ surj(X,Y)\ have "X = (\y\Y. {x\X . F`x = y})" using apply_type by fastforce show ?thesis proof (cases "Finite(Y)") case True with \X = (\y\Y. {x\X . F`x = y})\ and assms show ?thesis using Finite_RepFun[THEN [2] Finite_Union, of Y "\y. {x\X . F`x = y}"] by auto next case False moreover from this have "Y \ |Y|" using cardinal_eqpoll eqpoll_sym eqpoll_imp_lepoll by simp moreover note assms moreover from calculation have "y \ Y \ |{x\X . F`x = y}| \ |Y|" for y using Infinite_imp_nats_lepoll[THEN lepoll_imp_cardinal_le, of Y "|{x\X . F`x = y}|"] cardinal_idem by auto ultimately have "|\y\Y. {x\X . F`x = y}| \ |Y|" using leqpoll_imp_cardinal_UN_le[of "|Y|" Y] Infinite_InfCard_cardinal[of Y] by simp moreover from \F \ Finite_to_one(X,Y) \ surj(X,Y)\ have "|Y| \ |X|" using surj_implies_cardinal_le by auto moreover note \X = (\y\Y. {x\X . F`x = y})\ ultimately show ?thesis using le_anti_sym by auto qed qed lemma cardinal_map_Un: assumes "Infinite(X)" "Finite(b)" shows "|{a \ b . a \ X}| = |X|" proof - have "(\a\X. a \ b) \ Finite_to_one(X,{a \ b . a \ X})" "(\a\X. a \ b) \ surj(X,{a \ b . a \ X})" unfolding surj_def proof fix d have "Finite({a \ X . a \ b = d})" (is "Finite(?Y(b,d))") using \Finite(b)\ proof (induct arbitrary:d) case 0 have "{a \ X . a \ 0 = d} = (if d\X then {d} else 0)" by auto then show ?case by simp next case (cons c b) from \c \ b\ have "?Y(cons(c,b),d) \ (if c\d then ?Y(b,d) \ ?Y(b,d-{c}) else 0)" by auto with cons show ?case using subset_Finite by simp qed moreover assume "d \ {x \ b . x \ X}" ultimately show "Finite({a \ X . (\x\X. x \ b) ` a = d})" by simp qed (auto intro:lam_funtype) with assms show ?thesis using Finite_to_one_surj_imp_cardinal_eq by fast qed end \ No newline at end of file diff --git a/thys/Delta_System_Lemma/Cofinality.thy b/thys/Delta_System_Lemma/Cofinality.thy --- a/thys/Delta_System_Lemma/Cofinality.thy +++ b/thys/Delta_System_Lemma/Cofinality.thy @@ -1,1205 +1,1205 @@ section\Cofinality\label{sec:cofinality}\ theory Cofinality imports ZF_Library begin subsection\Basic results and definitions\ text\A set \<^term>\X\ is \<^emph>\cofinal\ in \<^term>\A\ (with respect to the relation - \<^term>\r\) if every element of \<^term>\A\ is “bounded - above” by some element of \<^term>\X\. Note that \<^term>\X\ does not need + \<^term>\r\) if every element of \<^term>\A\ is ``bounded + above'' by some element of \<^term>\X\. Note that \<^term>\X\ does not need to be a subset of \<^term>\A\.\ definition cofinal :: "[i,i,i] \ o" where "cofinal(X,A,r) \ \a\A. \x\X. \a,x\\r \ a = x" (* (* Alternative definitions *) definition cofinal_predic :: "[i,i,[i,i]\o] \ o" where "cofinal_predic(X,A,r) \ \a\A. \x\X. r(a,x) \ a = x" definition f_cofinal :: "[i\i,i,i,i] \ o" where "f_cofinal(f,C,A,r) \ \a\A. \x\C. \a,f(x)\\r \ a = f(x)" (* The next definition doesn't work if 0 is the top element of A. But it works for limit ordinals. *) definition cofinal_fun' :: "[i,i,i] \ o" where "cofinal_fun'(f,A,r) == f_cofinal(\x. f`x,domain(f),A, r)" *) text\A function is cofinal if it range is.\ definition cofinal_fun :: "[i,i,i] \ o" where "cofinal_fun(f,A,r) \ \a\A. \x\domain(f). \a,f`x\\r \ a = f`x" lemma cofinal_funI: assumes "\a. a\A \ \x\domain(f). \a,f`x\\r \ a = f`x" shows "cofinal_fun(f,A,r)" using assms unfolding cofinal_fun_def by simp lemma cofinal_funD: assumes "cofinal_fun(f,A,r)" "a\A" shows "\x\domain(f). \a,f`x\\r \ a = f`x" using assms unfolding cofinal_fun_def by simp lemma cofinal_in_cofinal: assumes "trans(r)" "cofinal(Y,X,r)" "cofinal(X,A,r)" shows "cofinal(Y,A,r)" unfolding cofinal_def proof fix a assume "a\A" moreover from \cofinal(X,A,r)\ have "b\A\\x\X. \b,x\\r \ b=x" for b unfolding cofinal_def by simp ultimately obtain y where "y\X" "\a,y\\r \ a=y" by auto moreover from \cofinal(Y,X,r)\ have "c\X\\y\Y. \c,y\\r \ c=y" for c unfolding cofinal_def by simp ultimately obtain x where "x\Y" "\y,x\\r \ y=x" by auto with \a\A\ \y\X\ \\a,y\\r \ a=y\ \trans(r)\ show "\x\Y. \a,x\\r \ a=x" unfolding trans_def by auto qed lemma codomain_is_cofinal: assumes "cofinal_fun(f,A,r)" "f:C \ D" shows "cofinal(D,A,r)" unfolding cofinal_def proof fix b assume "b \ A" moreover from assms have "a\A \ \x\domain(f). \a, f ` x\ \ r \ a = f`x" for a unfolding cofinal_fun_def by simp ultimately obtain x where "x\domain(f)" "\b, f ` x\ \ r \ b = f`x" by blast moreover from \f:C \ D\ \x\domain(f)\ have "f`x\D" using domain_of_fun apply_rangeI by simp ultimately show "\y\D. \b, y\ \ r \ b = y" by auto qed lemma cofinal_range_iff_cofinal_fun: assumes "function(f)" shows "cofinal(range(f),A,r) \ cofinal_fun(f,A,r)" unfolding cofinal_fun_def proof (intro iffI ballI) fix a assume "a\A" \cofinal(range(f),A,r)\ then obtain y where "y\range(f)" "\a,y\ \ r \ a = y" unfolding cofinal_def by blast moreover from this obtain x where "\x,y\\f" unfolding range_def domain_def converse_def by blast moreover note \function(f)\ ultimately have "\a, f ` x\ \ r \ a = f ` x" using function_apply_equality by blast with \\x,y\\f\ show "\x\domain(f). \a, f ` x\ \ r \ a = f ` x" by blast next assume "\a\A. \x\domain(f). \a, f ` x\ \ r \ a = f ` x" with assms show "cofinal(range(f), A, r)" using function_apply_Pair[of f] unfolding cofinal_def by fast qed lemma cofinal_comp: assumes "f\ mono_map(C,s,D,r)" "cofinal_fun(f,D,r)" "h:B \ C" "cofinal_fun(h,C,s)" "trans(r)" shows "cofinal_fun(f O h,D,r)" unfolding cofinal_fun_def proof fix a from \f\ mono_map(C,s,D,r)\ have "f:C \ D" using mono_map_is_fun by simp with \h:B \ C\ have "domain(f) = C" "domain(h) = B" using domain_of_fun by simp_all moreover assume "a \ D" moreover note \cofinal_fun(f,D,r)\ ultimately obtain c where "c\C" "\a, f ` c\ \ r \ a = f ` c" unfolding cofinal_fun_def by blast with \cofinal_fun(h,C,s)\ \domain(h) = B\ obtain b where "b \ B" "\c, h ` b\ \ s \ c = h ` b" unfolding cofinal_fun_def by blast moreover from this and \h:B \ C\ have "h`b \ C" by simp moreover note \f \ mono_map(C,s,D,r)\ \c\C\ ultimately have "\f`c, f` (h ` b)\ \ r \ f`c = f` (h ` b)" unfolding mono_map_def by blast with \\a, f ` c\ \ r \ a = f ` c\ \trans(r)\ \h:B \ C\ \b\B\ have "\a, (f O h) ` b\ \ r \ a = (f O h) ` b" using transD by auto moreover from \h:B \ C\ \domain(f) = C\ \domain(h) = B\ have "domain(f O h) = B" using range_fun_subset_codomain by blast moreover note \b\B\ ultimately show "\x\domain(f O h). \a, (f O h) ` x\ \ r \ a = (f O h) ` x" by blast qed definition cf_fun :: "[i,i] \ o" where "cf_fun(f,\) \ cofinal_fun(f,\,Memrel(\))" lemma cf_funI[intro!]: "cofinal_fun(f,\,Memrel(\)) \ cf_fun(f,\)" unfolding cf_fun_def by simp lemma cf_funD[dest!]: "cf_fun(f,\) \ cofinal_fun(f,\,Memrel(\))" unfolding cf_fun_def by simp lemma cf_fun_comp: assumes "Ord(\)" "f\ mono_map(C,s,\,Memrel(\))" "cf_fun(f,\)" "h:B \ C" "cofinal_fun(h,C,s)" shows "cf_fun(f O h,\)" using assms cofinal_comp[OF _ _ _ _ trans_Memrel] by auto definition cf :: "i\i" where "cf(\) \ \ \. \A. A\\ \ cofinal(A,\,Memrel(\)) \ \ = ordertype(A,Memrel(\))" lemma Ord_cf [TC]: "Ord(cf(\))" unfolding cf_def using Ord_Least by simp lemma gamma_cofinal_gamma: assumes "Ord(\)" shows "cofinal(\,\,Memrel(\))" unfolding cofinal_def by auto lemma cf_is_ordertype: assumes "Ord(\)" shows "\A. A\\ \ cofinal(A,\,Memrel(\)) \ cf(\) = ordertype(A,Memrel(\))" (is "?P(cf(\))") using gamma_cofinal_gamma LeastI[of ?P \] ordertype_Memrel[symmetric] assms unfolding cf_def by blast lemma cf_fun_succ': assumes "Ord(\)" "Ord(\)" "f:\\succ(\)" shows "(\x\\. f`x=\) \ cf_fun(f,succ(\))" proof (intro iffI) assume "(\x\\. f`x=\)" with assms show "cf_fun(f,succ(\))" using domain_of_fun[OF \f:\\succ(\)\] unfolding cf_fun_def cofinal_fun_def by auto next assume "cf_fun(f,succ(\))" with assms obtain x where "x\\" "\\,f`x\ \ Memrel(succ(\)) \ \ = f ` x" using domain_of_fun[OF \f:\\succ(\)\] unfolding cf_fun_def cofinal_fun_def by auto moreover from \Ord(\)\ have "\\,y\ \ Memrel(succ(\))" for y using foundation unfolding Memrel_def by blast ultimately show "\x\\. f ` x = \" by blast qed lemma cf_fun_succ: "Ord(\) \ f:1\succ(\) \ f`0=\ \ cf_fun(f,succ(\))" using cf_fun_succ' by blast lemma ordertype_0_not_cofinal_succ: assumes "ordertype(A,Memrel(succ(i))) = 0" "A\succ(i)" "Ord(i)" shows "\cofinal(A,succ(i),Memrel(succ(i)))" proof have 1:"ordertype(A,Memrel(succ(i))) = ordertype(0,Memrel(0))" using \ordertype(A,Memrel(succ(i))) = 0\ ordertype_0 by simp from \A\succ(i)\ \Ord(i)\ have "\f. f \ \A, Memrel(succ(i))\ \ \0, Memrel(0)\" using well_ord_Memrel well_ord_subset ordertype_eq_imp_ord_iso[OF 1] Ord_0 by blast then have "A=0" using ord_iso_is_bij bij_imp_eqpoll eqpoll_0_is_0 by blast moreover assume "cofinal(A, succ(i), Memrel(succ(i)))" moreover note \Ord(i)\ ultimately show "False" using not_mem_empty unfolding cofinal_def by auto qed text\I thank Edwin Pacheco Rodríguez for the following lemma.\ lemma cf_succ: assumes "Ord(\)" shows "cf(succ(\)) = 1" proof - define f where "f \ {\0,\\}" then have "f : 1\succ(\)" "f`0 = \" using fun_extend3[of 0 0 "succ(\)" 0 \] singleton_0 by auto with assms have "cf_fun(f,succ(\))" using cf_fun_succ unfolding cofinal_fun_def by simp from \f:1\succ(\)\ have "0\domain(f)" using domain_of_fun by simp define A where "A={f`0}" with \cf_fun(f,succ(\))\ \0\domain(f)\ \f`0=\\ have "cofinal(A,succ(\),Memrel(succ(\)))" unfolding cofinal_def cofinal_fun_def by simp moreover from \f`0=\\ \A={f`0}\ have "A \ succ(\)" unfolding succ_def by auto moreover from \Ord(\)\ \A\ succ(\)\ have "well_ord(A,Memrel(succ(\)))" using Ord_succ well_ord_Memrel well_ord_subset relation_Memrel by blast moreover from \Ord(\)\ have "\(\A. A \ succ(\) \ cofinal(A, succ(\), Memrel(succ(\))) \ 0 = ordertype(A, Memrel(succ(\))))" (is "\?P(0)") using ordertype_0_not_cofinal_succ unfolding cf_def by auto moreover have "1 = ordertype(A,Memrel(succ(\)))" proof - from \A={f`0}\ have "A\1" using singleton_eqpoll_1 by simp with \well_ord(A,Memrel(succ(\)))\ show ?thesis using nat_1I ordertype_eq_n by simp qed ultimately show "cf(succ(\)) = 1" using Ord_1 Least_equality[of ?P 1] unfolding cf_def by blast qed lemma cf_zero [simp]: "cf(0) = 0" unfolding cf_def cofinal_def using ordertype_0 subset_empty_iff Least_le[of _ 0] by auto lemma surj_is_cofinal: "f \ surj(\,\) \ cf_fun(f,\)" unfolding surj_def cofinal_fun_def cf_fun_def using domain_of_fun by force lemma cf_zero_iff: "Ord(\) \ cf(\) = 0 \ \ = 0" proof (intro iffI) assume "\ = 0" "Ord(\)" then show "cf(\) = 0" using cf_zero by simp next assume "cf(\) = 0" "Ord(\)" moreover from this obtain A where "A\\" "cf(\) = ordertype(A,Memrel(\))" "cofinal(A,\,Memrel(\))" using cf_is_ordertype by blast ultimately have "cofinal(0,\,Memrel(\))" using ordertype_zero_imp_zero[of A "Memrel(\)"] by simp then show "\=0" unfolding cofinal_def by blast qed \ \TODO: define Succ (predicate for successor ordinals)\ lemma cf_eq_one_iff: assumes "Ord(\)" shows "cf(\) = 1 \ (\\. Ord(\) \ \ = succ(\))" proof (intro iffI) assume "\\. Ord(\) \ \ = succ(\)" then show "cf(\) = 1" using cf_succ by auto next assume "cf(\) = 1" moreover from assms obtain A where "A \ \" "cf(\) = ordertype(A,Memrel(\))" "cofinal(A,\,Memrel(\))" using cf_is_ordertype by blast ultimately have "ordertype(A,Memrel(\)) = 1" by simp moreover define f where "f\converse(ordermap(A,Memrel(\)))" moreover from this \ordertype(A,Memrel(\)) = 1\ \A \ \\ assms have "f \ surj(1,A)" using well_ord_subset[OF well_ord_Memrel, THEN ordermap_bij, THEN bij_converse_bij, of \ A] bij_is_surj by simp with \cofinal(A,\,Memrel(\))\ have "\a\\. \a, f`0\ \ Memrel(\) \ a = f`0" unfolding cofinal_def surj_def by auto with assms \A \ \\ \f \ surj(1,A)\ show "\\. Ord(\) \ \ = succ(\)" using Ord_has_max_imp_succ[of \ "f`0"] surj_is_fun[of f 1 A] apply_type[of f 1 "\_.A" 0] unfolding lt_def by (auto intro:Ord_in_Ord) qed lemma ordertype_in_cf_imp_not_cofinal: assumes "ordertype(A,Memrel(\)) \ cf(\)" "A \ \" shows "\ cofinal(A,\,Memrel(\))" proof note \A \ \\ moreover assume "cofinal(A,\,Memrel(\))" ultimately have "\B. B \ \ \ cofinal(B, \, Memrel(\)) \ ordertype(A,Memrel(\)) = ordertype(B, Memrel(\))" (is "?P(ordertype(A,_))") by blast moreover from assms have "ordertype(A,Memrel(\)) < cf(\)" using Ord_cf ltI by blast ultimately show "False" unfolding cf_def using less_LeastE[of ?P "ordertype(A,Memrel(\))"] by auto qed lemma cofinal_mono_map_cf: assumes "Ord(\)" shows "\j \ mono_map(cf(\), Memrel(cf(\)), \, Memrel(\)) . cf_fun(j,\)" proof - note assms moreover from this obtain A where "A \ \" "cf(\) = ordertype(A,Memrel(\))" "cofinal(A,\,Memrel(\))" using cf_is_ordertype by blast moreover define j where "j \ converse(ordermap(A,Memrel(\)))" moreover from calculation have "j :cf(\) \\<^sub>< \" using ordertype_ord_iso[THEN ord_iso_sym, THEN ord_iso_is_mono_map, THEN mono_map_mono, of A "Memrel(\)" \] well_ord_Memrel[THEN well_ord_subset] by simp moreover from calculation have "j \ surj(cf(\),A)" using well_ord_Memrel[THEN well_ord_subset, THEN ordertype_ord_iso, THEN ord_iso_sym, of \ A, THEN ord_iso_is_bij, THEN bij_is_surj] by simp with \cofinal(A,\,Memrel(\))\ have "cf_fun(j,\)" using cofinal_range_iff_cofinal_fun[of j \ "Memrel(\)"] surj_range[of j "cf(\)" A] surj_is_fun fun_is_function by fastforce with \j \ mono_map(_,_,_,_)\ show ?thesis by auto qed subsection\The factorization lemma\ text\In this subsection we prove a factorization lemma for cofinal functions into ordinals, which shows that any cofinal function between ordinals can be -“decomposed” in such a way that a commutative triangle of strictly increasing +``decomposed'' in such a way that a commutative triangle of strictly increasing maps arises. The factorization lemma has a kind of fundamental character, in that the rest of the basic results on cofinality (for, instance, idempotence) follow easily from it, in a more algebraic way. This is a consequence that the proof encapsulates uses of transfinite recursion in the basic theory of cofinality; indeed, only one use is needed. In the setting of Isabelle/ZF, this is convenient since the machinery of recursion is pretty clumsy. On the downside, this way of presenting things results in a longer proof of the factorization lemma. This approach was taken by the author in the notes \cite{apunte_st} for an introductory course in Set Theory. To organize the use of the hypotheses of the factorization lemma, we set up a locale containing all the relevant ingredients. \ locale cofinal_factor = fixes j \ \ \ f assumes j_mono: "j :\ \\<^sub>< \" and ords: "Ord(\)" "Ord(\)" "Limit(\)" and f_type: "f: \ \ \" begin text\Here, \<^term>\f\ is cofinal function from \<^term>\\\ to \<^term>\\\, and the ordinal \<^term>\\\ is meant to be the cofinality of \<^term>\\\. Hence, there exists an increasing map \<^term>\j\ from \<^term>\\\ to \<^term>\\\ by the last lemma. The main goal is to construct an increasing function \<^term>\g:\\\\ such that the composition \<^term>\f O g\ is still cofinal but also increasing.\ definition factor_body :: "[i,i,i] \ o" where "factor_body(\,h,x) \ (x\\ \ j`\ \ f`x \ (\\<\ . f`(h`\) < f`x)) \ x=\" definition factor_rec :: "[i,i] \ i" where "factor_rec(\,h) \ \ x. factor_body(\,h,x)" txt\\<^term>\factor_rec\ is the inductive step for the definition by transfinite recursion of the \<^emph>\factor\ function (called \<^term>\g\ above), which in turn is obtained by minimizing the predicate \<^term>\factor_body\. Next we show that this predicate is monotonous.\ lemma factor_body_mono: assumes "\\\" "\<\" "factor_body(\,\x\\. G(x),x)" shows "factor_body(\,\x\\. G(x),x)" proof - from \\<\\ have "\\\" using ltD by simp moreover note \\\\\ moreover from calculation have "\\\" using ords ltD Ord_cf Ord_trans by blast ultimately have "j`\ \ j`\" using j_mono mono_map_increasing by blast moreover from \\\\\ have "j`\\\" using j_mono domain_of_fun apply_rangeI mono_map_is_fun by force moreover from this have "Ord(j`\)" using Ord_in_Ord ords Limit_is_Ord by auto ultimately have "j`\ \ j`\" unfolding lt_def by blast then have "j`\ \ f`\ \ j`\ \ f`\" for \ using le_trans by blast moreover have "f`((\w\\. G(w))`y) < f`z" if "z\\" "\x<\. f`((\w\\. G(w))`x) < f`z" "y<\" for y z proof - note \y<\\ also note \\<\\ finally have "y<\" by simp with \\x<\. f`((\w\\. G(w))`x) < f`z\ have "f ` ((\w\\. G(w)) ` y) < f ` z" by simp moreover from \y<\\ \y<\\ have "(\w\\. G(w)) ` y = (\w\\. G(w)) ` y" using beta_if by (auto dest:ltD) ultimately show ?thesis by simp qed moreover note \factor_body(\,\x\\. G(x),x)\ ultimately show ?thesis unfolding factor_body_def by blast qed lemma factor_body_simp[simp]: "factor_body(\,g,\)" unfolding factor_body_def by simp lemma factor_rec_mono: assumes "\\\" "\<\" shows "factor_rec(\,\x\\. G(x)) \ factor_rec(\,\x\\. G(x))" unfolding factor_rec_def using assms ords factor_body_mono Least_antitone by simp text\We now define the factor as higher-order function. Later it will be restricted to a set to obtain a bona fide function of type @{typ i}.\ definition factor :: "i \ i" where "factor(\) \ transrec(\,factor_rec)" lemma factor_unfold: "factor(\) = factor_rec(\,\x\\. factor(x))" using def_transrec[OF factor_def] . lemma factor_mono: assumes "\\\" "\<\" "factor(\)\\" "factor(\)\\" shows "factor(\) \ factor(\)" proof - have "factor(\) = factor_rec(\, \x\\. factor(x))" using factor_unfold . also from assms and factor_rec_mono have "... \ factor_rec(\, \x\\. factor(x))" by simp also have "factor_rec(\, \x\\. factor(x)) = factor(\)" using def_transrec[OF factor_def, symmetric] . finally show ?thesis . qed text\The factor satisfies the predicate body of the minimization.\ lemma factor_body_factor: "factor_body(\,\x\\. factor(x),factor(\))" using ords factor_unfold[of \] LeastI[of "factor_body(_,_)" \] unfolding factor_rec_def by simp lemma factor_type [TC]: "Ord(factor(\))" using ords factor_unfold[of \] unfolding factor_rec_def by simp text\The value \<^term>\\\ in \<^term>\factor_body\ (and therefore, in -\<^term>\factor\) is meant to be a “default value”. Whenever it is not +\<^term>\factor\) is meant to be a ``default value''. Whenever it is not attained, the factor function behaves as expected: It is increasing and its composition with \<^term>\f\ also is.\ lemma f_factor_increasing: assumes "\\\" "\<\" "factor(\)\\" shows "f`factor(\) < f`factor(\)" proof - from assms have "f ` ((\x\\. factor(x)) ` \) < f ` factor(\)" using factor_unfold[of \] ords LeastI[of "factor_body(\,\x\\. factor(x))"] unfolding factor_rec_def factor_body_def by (auto simp del:beta_if) with \\<\\ show ?thesis using ltD by auto qed lemma factor_increasing: assumes "\\\" "\<\" "factor(\)\\" "factor(\)\\" shows "factor(\))" using assms f_factor_increasing factor_mono by (force intro:le_neq_imp_lt) lemma factor_in_delta: assumes "factor(\) \ \" shows "factor(\) \ \" using assms factor_body_factor ords unfolding factor_body_def by auto text\Finally, we define the (set) factor function as the restriction of factor to the ordinal \<^term>\\\.\ definition fun_factor :: "i" where "fun_factor \ \\\\. factor(\)" lemma fun_factor_is_mono_map: assumes "\\. \ \ \ \ factor(\) \ \" shows "fun_factor \ mono_map(\, Memrel(\), \, Memrel(\))" unfolding mono_map_def proof (intro CollectI ballI impI) (* Proof that \<^term>\fun_factor\ respects membership *) fix \ \ assume "\\\" "\\\" moreover note assms moreover from calculation have "factor(\)\\" "factor(\)\\" "Ord(\)" using factor_in_delta Ord_in_Ord ords by auto moreover assume "\\, \\ \ Memrel(\)" ultimately show "\fun_factor ` \, fun_factor ` \\ \ Memrel(\)" unfolding fun_factor_def using ltI factor_increasing[THEN ltD] factor_in_delta by simp next (* Proof of type *) from assms show "fun_factor : \ \ \" unfolding fun_factor_def using ltI lam_type factor_in_delta by simp qed lemma f_fun_factor_is_mono_map: assumes "\\. \ \ \ \ factor(\) \ \" shows "f O fun_factor \ mono_map(\, Memrel(\), \, Memrel(\))" unfolding mono_map_def using f_type proof (intro CollectI ballI impI comp_fun[of _ _ \]) from assms show "fun_factor : \ \ \" using fun_factor_is_mono_map mono_map_is_fun by simp (* Proof that f O ?g respects membership *) fix \ \ assume "\\, \\ \ Memrel(\)" then have "\<\" using Ord_in_Ord[of "\"] ltI ords by blast assume "\\\" "\\\" moreover from this and assms have "factor(\)\\" "factor(\)\\" by auto moreover have "Ord(\)" "\\0" using ords Limit_is_Ord by auto moreover note \\<\\ \fun_factor : \ \ \\ ultimately show "\(f O fun_factor) ` \, (f O fun_factor) ` \\ \ Memrel(\)" using ltD[of "f ` factor(\)" "f ` factor(\)"] f_factor_increasing apply_in_range f_type unfolding fun_factor_def by auto qed end (* cofinal_factor *) text\We state next the factorization lemma.\ lemma cofinal_fun_factorization: notes le_imp_subset [dest] lt_trans2 [trans] assumes "Ord(\)" "Limit(\)" "f: \ \ \" "cf_fun(f,\)" shows "\g \ cf(\) \\<^sub>< \. f O g : cf(\) \\<^sub>< \ \ cofinal_fun(f O g,\,Memrel(\))" proof - from \Limit(\)\ have "Ord(\)" using Limit_is_Ord by simp then obtain j where "j :cf(\) \\<^sub>< \" "cf_fun(j,\)" using cofinal_mono_map_cf by blast then have "domain(j) = cf(\)" using domain_of_fun mono_map_is_fun by force from \j \ _\ assms interpret cofinal_factor j \ "cf(\)" by (unfold_locales) (simp_all) text\The core of the argument is to show that the factor function indeed maps into \<^term>\\\, therefore its values satisfy the first disjunct of \<^term>\factor_body\. This holds in turn because no restriction of the factor composed with \<^term>\f\ to a proper initial segment of \<^term>\cf(\)\ can be cofinal in \<^term>\\\ by definition of cofinality. Hence there must be a witness that satisfies the first disjunct.\ have factor_not_delta: "factor(\)\\" if "\ \ cf(\)" for \ text\For this, we induct on \<^term>\\\ ranging over \<^term>\cf(\)\.\ proof (induct \ rule:Ord_induct[OF _ Ord_cf[of \]]) case 1 with that show ?case . next case (2 \) then have IH: "z\\ \ factor(z)\\" for z by simp define h where "h \ \x\\. f`factor(x)" from IH have "z\\ \ factor(z) \ \" for z using factor_in_delta by blast with \f:\\\\ have "h : \ \ \" unfolding h_def using apply_funtype lam_type by auto then have "h : \ \\<^sub>< \" unfolding mono_map_def proof (intro CollectI ballI impI) fix x y assume "x\\" "y\\" moreover from this and IH have "factor(y) \ \" by simp moreover from calculation and \h \ \ \ \\ have "h`x \ \" "h`y \ \" by simp_all moreover from \\\cf(\)\ and \y\\\ have "y \ cf(\)" using Ord_trans Ord_cf by blast moreover from this have "Ord(y)" using Ord_cf Ord_in_Ord by blast moreover assume "\x,y\ \ Memrel(\)" moreover from calculation have "xh ` x, h ` y\ \ Memrel(\)" unfolding h_def using f_factor_increasing ltD by (auto) qed with \\\cf(\)\ \Ord(\)\ have "ordertype(h``\,Memrel(\)) = \" (* Maybe should use range(h) *) using mono_map_ordertype_image[of \] Ord_cf Ord_in_Ord by blast also note \\ \cf(\)\ finally have "ordertype(h``\,Memrel(\)) \ cf(\)" by simp moreover from \h \ \ \ \\ have "h``\ \ \" using mono_map_is_fun Image_sub_codomain by blast ultimately have "\ cofinal(h``\,\,Memrel(\))" using ordertype_in_cf_imp_not_cofinal by simp then obtain \_0 where "\_0\\" "\x\h `` \. \ \\_0, x\ \ Memrel(\) \ \_0 \ x" unfolding cofinal_def by auto with \Ord(\)\ \h``\ \ \\ have "\x\h `` \. x \ \_0" using well_ord_Memrel[of \] well_ord_is_linear[of \ "Memrel(\)"] unfolding linear_def by blast from \\_0 \ \\ \j \ mono_map(_,_,\,_)\ \Ord(\)\ have "j`\ \ \" using mono_map_is_fun apply_in_range by force with \\_0 \ \\ \Ord(\)\ have "\_0 \ j`\ \ \" using Un_least_mem_iff Ord_in_Ord by auto with \cf_fun(f,\)\ obtain \ where "\\domain(f)" "\\_0 \ j`\, f ` \\ \ Memrel(\) \ \_0 \ j`\ = f ` \" by (auto simp add:cofinal_fun_def) blast moreover from this and \f:\\\\ have "\ \ \" using domain_of_fun by auto moreover note \Ord(\)\ moreover from this and \f:\\\\ \\_0 \ \\ have "Ord(f`\)" using apply_in_range Ord_in_Ord by blast moreover from calculation and \\_0 \ \\ and \Ord(\)\ and \j`\ \ \\ have "Ord(\_0)" "Ord(j`\)" "Ord(\)" using Ord_in_Ord by auto moreover from \\x\h `` \. x \ \_0\ \Ord(\_0)\ \h:\\\\ have "x\\ \ h`x < \_0" for x using fun_is_function[of h \ "\_. \"] Image_subset_Ord_imp_lt domain_of_fun[of h \ "\_. \"] by blast moreover have "x\\ \ h`x < f`\" for x proof - fix x assume "x\\" with \\x\h `` \. x \ \_0\ \Ord(\_0)\ \h:\\\\ have "h`x < \_0" using fun_is_function[of h \ "\_. \"] Image_subset_Ord_imp_lt domain_of_fun[of h \ "\_. \"] by blast also from \\\_0 \ _, f ` \\ \ Memrel(\) \ \_0 \ _= f ` \\ \Ord(f`\)\ \Ord(\_0)\ \Ord(j`\)\ have "\_0 \ f`\" using Un_leD1[OF leI [OF ltI]] Un_leD1[OF le_eqI] by blast finally show "h`x < f`\" . qed ultimately have "factor_body(\,\x\\. factor(x),\)" unfolding h_def factor_body_def using ltD by (auto dest:Un_memD2 Un_leD2[OF le_eqI]) with \Ord(\)\ have "factor(\) \ \" using factor_unfold[of \] Least_le unfolding factor_rec_def by auto with \\\\\ \Ord(\)\ have "factor(\) \ \" using leI[of \] ltI[of \] by (auto dest:ltD) then show ?case by (auto elim:mem_irrefl) qed moreover have "cofinal_fun(f O fun_factor, \, Memrel(\))" proof (intro cofinal_funI) fix a assume "a \ \" with \cf_fun(j,\)\ \domain(j) = cf(\)\ obtain x where "x\cf(\)" "a \ j`x \ a = j`x" by (auto simp add:cofinal_fun_def) blast with factor_not_delta have "x \ domain(f O fun_factor)" using f_fun_factor_is_mono_map mono_map_is_fun domain_of_fun by force moreover have "a \ (f O fun_factor) `x \ a = (f O fun_factor) `x" proof - from \x\cf(\)\ factor_not_delta have "j ` x \ f ` factor(x)" using mem_not_refl factor_body_factor factor_in_delta unfolding factor_body_def by auto with \a \ j`x \ a = j`x\ have "a \ f ` factor(x) \ a = f ` factor(x)" using ltD by blast with \x\cf(\)\ show ?thesis using lam_funtype[of "cf(\)" factor] unfolding fun_factor_def by auto qed moreover note \a \ \\ moreover from calculation and \Ord(\)\ and factor_not_delta have "(f O fun_factor) `x \ \" using Limit_nonzero apply_in_range mono_map_is_fun[of "f O fun_factor"] f_fun_factor_is_mono_map by blast ultimately show "\x \ domain(f O fun_factor). \a, (f O fun_factor) ` x\ \ Memrel(\) \ a = (f O fun_factor) `x" by blast qed ultimately show ?thesis using fun_factor_is_mono_map f_fun_factor_is_mono_map by blast qed text\As a final observation in this part, we note that if the original cofinal map was increasing, then the factor function is also cofinal.\ lemma factor_is_cofinal: assumes "Ord(\)" "Ord(\)" "f :\ \\<^sub>< \" "f O g \ mono_map(\,r,\,Memrel(\))" "cofinal_fun(f O g,\,Memrel(\))" "g: \ \ \" shows "cf_fun(g,\)" unfolding cf_fun_def cofinal_fun_def proof fix a assume "a \ \" with \f \ mono_map(\,_,\,_)\ have "f`a \ \" using mono_map_is_fun by force with \cofinal_fun(f O g,\,_)\ obtain y where "y\\" "\f`a, (f O g) ` y\ \ Memrel(\) \ f`a = (f O g) ` y" unfolding cofinal_fun_def using domain_of_fun[OF \g:\ \ \\] by blast with \g:\ \ \\ have "\f`a, f ` (g ` y)\ \ Memrel(\) \ f`a = f ` (g ` y)" "g`y \ \" using comp_fun_apply[of g \ \ y f] by auto with assms(1-3) and \a\\\ have "\a, g ` y\ \ Memrel(\) \ a = g ` y" using Memrel_mono_map_reflects Memrel_mono_map_is_inj[of \ f \ \] inj_apply_equality[of f \ \] by blast with \y\\\ show "\x\domain(g). \a, g ` x\ \ Memrel(\) \ a = g ` x" using domain_of_fun[OF \g:\ \ \\] by blast qed subsection\Classical results on cofinalities\ text\Now the rest of the results follow in a more algebraic way. The next proof one invokes a case analysis on whether the argument is zero, a successor ordinal or a limit one; the last case being the most relevant one and is immediate from the factorization lemma.\ lemma cf_le_domain_cofinal_fun: assumes "Ord(\)" "Ord(\)" "f:\ \ \" "cf_fun(f,\)" shows "cf(\)\\" using assms proof (cases rule:Ord_cases) case 0 with \Ord(\)\ show ?thesis using Ord_0_le by simp next case (succ \) with assms obtain x where "x\\" "f`x=\" using cf_fun_succ' by blast then have "\\0" by blast let ?f="{\0,f`x\}" from \f`x=\\ have "?f:1\succ(\)" using singleton_0 singleton_fun[of 0 \] singleton_subsetI fun_weaken_type by simp with \Ord(\)\ \f`x=\\ have "cf(succ(\)) = 1" using cf_succ by simp with \\\0\ succ show ?thesis using Ord_0_lt_iff succ_leI \Ord(\)\ by simp next case (limit) with assms obtain g where "g :cf(\) \\<^sub>< \" using cofinal_fun_factorization by blast with assms show ?thesis using mono_map_imp_le by simp qed lemma cf_ordertype_cofinal: assumes "Limit(\)" "A\\" "cofinal(A,\,Memrel(\))" shows "cf(\) = cf(ordertype(A,Memrel(\)))" proof (intro le_anti_sym) from \Limit(\)\ have "Ord(\)" using Limit_is_Ord by simp with \A \ \\ have "well_ord(A,Memrel(\))" using well_ord_Memrel well_ord_subset by blast then obtain f \ where "f:\\, Memrel(\)\ \ \A,Memrel(\)\" "Ord(\)" "\ = ordertype(A,Memrel(\))" using ordertype_ord_iso Ord_ordertype ord_iso_sym by blast moreover from this have "f: \ \ A" using ord_iso_is_mono_map mono_map_is_fun[of f _ "Memrel(\)"] by blast moreover from this have "function(f)" using fun_is_function by simp moreover from \f:\\, Memrel(\)\ \ \A,Memrel(\)\\ have "range(f) = A" using ord_iso_is_bij bij_is_surj surj_range by blast moreover note \cofinal(A,\,_)\ ultimately have "cf_fun(f,\)" using cofinal_range_iff_cofinal_fun by blast moreover from \Ord(\)\ obtain h where "h :cf(\) \\<^sub>< \" "cf_fun(h,\)" using cofinal_mono_map_cf by blast moreover from \Ord(\)\ have "trans(Memrel(\))" using trans_Memrel by simp moreover note \A\\\ ultimately have "cofinal_fun(f O h,\,Memrel(\))" using cofinal_comp ord_iso_is_mono_map[OF \f:\\,_\ \ \A,_\\] mono_map_is_fun mono_map_mono by blast moreover from \f:\\A\ \A\\\ \h\mono_map(cf(\),_,\,_)\ have "f O h : cf(\) \ \" using Pi_mono[of A \] comp_fun mono_map_is_fun by blast moreover note \Ord(\)\ \Ord(\)\ \\ = ordertype(A,Memrel(\))\ ultimately show "cf(\) \ cf(ordertype(A,Memrel(\)))" using cf_le_domain_cofinal_fun[of _ _ "f O h"] by (auto simp add:cf_fun_def) (********************************************************) from \f:\\, _\ \ \A,_\\ \A\\\ have "f :\ \\<^sub>< \" using mono_map_mono[OF ord_iso_is_mono_map] by simp then have "f: \ \ \" using mono_map_is_fun by simp with \cf_fun(f,\)\ \Limit(\)\ \Ord(\)\ obtain g where "g :cf(\) \\<^sub>< \" "f O g :cf(\) \\<^sub>< \" "cofinal_fun(f O g,\,Memrel(\))" using cofinal_fun_factorization by blast moreover from this have "g:cf(\)\\" using mono_map_is_fun by simp moreover note \Ord(\)\ moreover from calculation and \f :\ \\<^sub>< \\ \Ord(\)\ have "cf_fun(g,\)" using factor_is_cofinal by blast moreover note \\ = ordertype(A,Memrel(\))\ ultimately show "cf(ordertype(A,Memrel(\))) \ cf(\)" using cf_le_domain_cofinal_fun[OF _ Ord_cf mono_map_is_fun] by simp qed lemma cf_idemp: assumes "Limit(\)" shows "cf(\) = cf(cf(\))" proof - from assms obtain A where "A\\" "cofinal(A,\,Memrel(\))" "cf(\) = ordertype(A,Memrel(\))" using Limit_is_Ord cf_is_ordertype by blast with assms have "cf(\) = cf(ordertype(A,Memrel(\)))" using cf_ordertype_cofinal by simp also have "... = cf(cf(\))" using \cf(\) = ordertype(A,Memrel(\))\ by simp finally show "cf(\) = cf(cf(\))" . qed lemma cf_le_cardinal: assumes "Limit(\)" shows "cf(\) \ |\|" proof - from assms have \Ord(\)\ using Limit_is_Ord by simp then obtain f where "f \ surj(|\|,\)" using Ord_cardinal_eqpoll unfolding eqpoll_def bij_def by blast with \Ord(\)\ show ?thesis using Card_is_Ord[OF Card_cardinal] surj_is_cofinal cf_le_domain_cofinal_fun[of \] surj_is_fun by blast qed lemma regular_is_Card: notes le_imp_subset [dest] assumes "Limit(\)" "\ = cf(\)" shows "Card(\)" proof - from assms have "|\| \ \" using Limit_is_Ord Ord_cardinal_le by blast also from \\ = cf(\)\ have "\ \ cf(\)" by simp finally have "|\| \ cf(\)" . with assms show ?thesis unfolding Card_def using cf_le_cardinal by force qed lemma Limit_cf: assumes "Limit(\)" shows "Limit(cf(\))" using Ord_cf[of \, THEN Ord_cases] \ \\<^term>\cf(\)\ being 0 or successor leads to contradiction\ proof (cases) case 1 with \Limit(\)\ show ?thesis using cf_zero_iff Limit_is_Ord by simp next case (2 \) moreover note \Limit(\)\ moreover from calculation have "cf(\) = 1" using cf_idemp cf_succ by fastforce ultimately show ?thesis using succ_LimitE cf_eq_one_iff Limit_is_Ord by auto qed lemma InfCard_cf: "Limit(\) \ InfCard(cf(\))" using regular_is_Card cf_idemp Limit_cf nat_le_Limit Limit_cf unfolding InfCard_def by simp lemma cf_le_cf_fun: notes [dest] = Limit_is_Ord assumes "cf(\) \ \" "Limit(\)" shows "\f. f:\ \ \ \ cf_fun(f, \)" proof - note assms moreover from this obtain h where h_cofinal_mono: "cf_fun(h,\)" "h :cf(\) \\<^sub>< \" "h : cf(\) \ \" using cofinal_mono_map_cf mono_map_is_fun by force moreover from calculation obtain g where "g \ inj(cf(\), \)" using le_imp_lepoll by blast from this and calculation(2,3,5) obtain f where "f \ surj(\, cf(\))" "f: \ \ cf(\)" using inj_imp_surj[OF _ Limit_has_0[THEN ltD]] surj_is_fun Limit_cf by blast moreover from this have "cf_fun(f,cf(\))" using surj_is_cofinal by simp moreover note h_cofinal_mono \Limit(\)\ moreover from calculation have "cf_fun(h O f,\)" using cf_fun_comp by blast moreover from calculation have "h O f \ \ -> \" using comp_fun by simp ultimately show ?thesis by blast qed lemma Limit_cofinal_fun_lt: notes [dest] = Limit_is_Ord assumes "Limit(\)" "f: \ \ \" "cf_fun(f,\)" "n\\" shows "\\\\. n < f`\" proof - from \Limit(\)\ \n\\\ have "succ(n) \ \" using Limit_has_succ[OF _ ltI, THEN ltD] by auto moreover note \f: \ \ _\ moreover from this have "domain(f) = \" using domain_of_fun by simp moreover note \cf_fun(f,\)\ ultimately obtain \ where "\ \ \" "succ(n) \ f`\ \ succ(n) = f `\" using cf_funD[THEN cofinal_funD] by blast moreover from this consider (1) "succ(n) \ f`\" | (2) "succ(n) = f `\" by blast then have "n < f`\" proof (cases) case 1 moreover have "n \ succ(n)" by simp moreover note \Limit(\)\ \f: \ \ _\ \\ \ \\ moreover from this have "Ord(f ` \)" using apply_type[of f \ "\_. \", THEN [2] Ord_in_Ord] by blast ultimately show ?thesis using Ord_trans[of n "succ(n)" "f ` \"] ltI by blast next case 2 have "n \ f ` \" by (simp add:2[symmetric]) with \Limit(\)\ \f: \ \ _\ \\ \ \\ show ?thesis using ltI apply_type[of f \ "\_. \", THEN [2] Ord_in_Ord] by blast qed ultimately show ?thesis by blast qed context includes Ord_dests Aleph_dests Aleph_intros Aleph_mem_dests mono_map_rules begin text\We end this section by calculating the cofinality of Alephs, for the zero and limit case. The successor case depends on $\AC$.\ lemma cf_nat: "cf(\) = \" using Limit_nat[THEN InfCard_cf] cf_le_cardinal[of \] Card_nat[THEN Card_cardinal_eq] le_anti_sym unfolding InfCard_def by auto lemma cf_Aleph_zero: "cf(\\<^bsub>0\<^esub>) = \\<^bsub>0\<^esub>" using cf_nat unfolding Aleph_def by simp lemma cf_Aleph_Limit: assumes "Limit(\)" shows "cf(\\<^bsub>\\<^esub>) = cf(\)" proof - note \Limit(\)\ moreover from this have "(\x\\. \\<^bsub>x\<^esub>) : \ \ \\<^bsub>\\<^esub>" (is "?f : _ \ _") using lam_funtype[of _ Aleph] fun_weaken_type[of _ _ _ "\\<^bsub>\\<^esub>"] by blast moreover from \Limit(\)\ have "x \ y \ \\<^bsub>x\<^esub> \ \\<^bsub>y\<^esub>" if "x \ \" "y \ \" for x y using that Ord_in_Ord[of \] Ord_trans[of _ _ \] by blast ultimately have "?f \ mono_map(\,Memrel(\),\\<^bsub>\\<^esub>, Memrel(\\<^bsub>\\<^esub>))" by auto with \Limit(\)\ have "?f \ \\, Memrel(\)\ \ \?f``\, Memrel(\\<^bsub>\\<^esub>)\" using mono_map_imp_ord_iso_Memrel[of \ "\\<^bsub>\\<^esub>" ?f] Card_Aleph (* Already an intro rule, but need it explicitly *) by blast then have "converse(?f) \ \?f``\, Memrel(\\<^bsub>\\<^esub>)\ \ \\, Memrel(\)\" using ord_iso_sym by simp with \Limit(\)\ have "ordertype(?f``\, Memrel(\\<^bsub>\\<^esub>)) = \" using ordertype_eq[OF _ well_ord_Memrel] ordertype_Memrel by auto moreover from \Limit(\)\ have "cofinal(?f``\, \\<^bsub>\\<^esub>, Memrel(\\<^bsub>\\<^esub>))" unfolding cofinal_def proof (standard, intro ballI) fix a assume "a\\\<^bsub>\\<^esub>" "\\<^bsub>\\<^esub> = (\i<\. \\<^bsub>i\<^esub>)" moreover from this obtain i where "i<\" "a\\\<^bsub>i\<^esub>" by auto moreover from this and \Limit(\)\ have "Ord(i)" using ltD Ord_in_Ord by blast moreover from \Limit(\)\ and calculation have "succ(i) \ \" using ltD by auto moreover from this and \Ord(i)\ have "\\<^bsub>i\<^esub> < \\<^bsub>succ(i)\<^esub>" by (auto) ultimately have "\a,\\<^bsub>i\<^esub>\ \ Memrel(\\<^bsub>\\<^esub>)" using ltD by (auto dest:Aleph_increasing) moreover from \i<\\ have "\\<^bsub>i\<^esub> \ ?f``\" using ltD apply_in_image[OF \?f : _ \ _\] by auto ultimately show "\x\?f `` \. \a, x\ \ Memrel(\\<^bsub>\\<^esub>) \ a = x" by blast qed moreover note \?f: \ \ \\<^bsub>\\<^esub>\ \Limit(\)\ ultimately show "cf(\\<^bsub>\\<^esub>) = cf(\)" using cf_ordertype_cofinal[OF Limit_Aleph Image_sub_codomain, of \ ?f \ \ ] Limit_is_Ord by simp qed end (* includes *) end \ No newline at end of file diff --git a/thys/Jacobson_Basic_Algebra/Group_Theory.thy b/thys/Jacobson_Basic_Algebra/Group_Theory.thy --- a/thys/Jacobson_Basic_Algebra/Group_Theory.thy +++ b/thys/Jacobson_Basic_Algebra/Group_Theory.thy @@ -1,1700 +1,1700 @@ (* Copyright (c) 2014-2019 by Clemens Ballarin This file is licensed under the 3-clause BSD license. *) theory Group_Theory imports Set_Theory begin hide_const monoid hide_const group hide_const inverse no_notation quotient (infixl "'/'/" 90) section \Monoids and Groups\ subsection \Monoids of Transformations and Abstract Monoids\ text \Def 1.1\ text \p 28, ll 28--30\ locale monoid = fixes M and composition (infixl "\" 70) and unit ("\") assumes composition_closed [intro, simp]: "\ a \ M; b \ M \ \ a \ b \ M" and unit_closed [intro, simp]: "\ \ M" and associative [intro]: "\ a \ M; b \ M; c \ M \ \ (a \ b) \ c = a \ (b \ c)" and left_unit [intro, simp]: "a \ M \ \ \ a = a" and right_unit [intro, simp]: "a \ M \ a \ \ = a" text \p 29, ll 27--28\ locale submonoid = monoid M "(\)" \ for N and M and composition (infixl "\" 70) and unit ("\") + assumes subset: "N \ M" and sub_composition_closed: "\ a \ N; b \ N \ \ a \ b \ N" and sub_unit_closed: "\ \ N" begin text \p 29, ll 27--28\ lemma sub [intro, simp]: "a \ N \ a \ M" using subset by blast text \p 29, ll 32--33\ sublocale sub: monoid N "(\)" \ by unfold_locales (auto simp: sub_composition_closed sub_unit_closed) end (* submonoid *) text \p 29, ll 33--34\ theorem submonoid_transitive: assumes "submonoid K N composition unit" and "submonoid N M composition unit" shows "submonoid K M composition unit" proof - interpret K: submonoid K N composition unit by fact interpret M: submonoid N M composition unit by fact show ?thesis by unfold_locales auto qed text \p 28, l 23\ locale transformations = fixes S :: "'a set" (* assumes non_vacuous: "S \ {}" *) (* Jacobson requires this but we don't need it, strange. *) text \Monoid of all transformations\ text \p 28, ll 23--24\ sublocale transformations \ monoid "S \\<^sub>E S" "compose S" "identity S" by unfold_locales (auto simp: PiE_def compose_eq compose_assoc Id_compose compose_Id) text \@{term N} is a monoid of transformations of the set @{term S}.\ text \p 29, ll 34--36\ locale transformation_monoid = transformations S + submonoid M "S \\<^sub>E S" "compose S" "identity S" for M and S begin text \p 29, ll 34--36\ lemma transformation_closed [intro, simp]: "\ \ \ M; x \ S \ \ \ x \ S" by (metis PiE_iff sub) text \p 29, ll 34--36\ lemma transformation_undefined [intro, simp]: "\ \ \ M; x \ S \ \ \ x = undefined" by (metis PiE_arb sub) end (* transformation_monoid *) subsection \Groups of Transformations and Abstract Groups\ context monoid begin text \Invertible elements\ text \p 31, ll 3--5\ definition invertible where "u \ M \ invertible u \ (\v \ M. u \ v = \ \ v \ u = \)" text \p 31, ll 3--5\ lemma invertibleI [intro]: "\ u \ v = \; v \ u = \; u \ M; v \ M \ \ invertible u" unfolding invertible_def by fast text \p 31, ll 3--5\ lemma invertibleE [elim]: "\ invertible u; \v. \ u \ v = \ \ v \ u = \; v \ M \ \ P; u \ M \ \ P" unfolding invertible_def by fast text \p 31, ll 6--7\ theorem inverse_unique: "\ u \ v' = \; v \ u = \; u \ M; v \ M; v' \ M \ \ v = v'" by (metis associative left_unit right_unit) text \p 31, l 7\ definition inverse where "inverse = (\u \ M. THE v. v \ M \ u \ v = \ \ v \ u = \)" text \p 31, l 7\ theorem inverse_equality: "\ u \ v = \; v \ u = \; u \ M; v \ M \ \ inverse u = v" unfolding inverse_def using inverse_unique by simp blast text \p 31, l 7\ lemma invertible_inverse_closed [intro, simp]: "\ invertible u; u \ M \ \ inverse u \ M" using inverse_equality by auto text \p 31, l 7\ lemma inverse_undefined [intro, simp]: "u \ M \ inverse u = undefined" by (simp add: inverse_def) text \p 31, l 7\ lemma invertible_left_inverse [simp]: "\ invertible u; u \ M \ \ inverse u \ u = \" using inverse_equality by auto text \p 31, l 7\ lemma invertible_right_inverse [simp]: "\ invertible u; u \ M \ \ u \ inverse u = \" using inverse_equality by auto text \p 31, l 7\ lemma invertible_left_cancel [simp]: "\ invertible x; x \ M; y \ M; z \ M \ \ x \ y = x \ z \ y = z" by (metis associative invertible_def left_unit) text \p 31, l 7\ lemma invertible_right_cancel [simp]: "\ invertible x; x \ M; y \ M; z \ M \ \ y \ x = z \ x \ y = z" by (metis associative invertible_def right_unit) text \p 31, l 7\ lemma inverse_unit [simp]: "inverse \ = \" using inverse_equality by blast text \p 31, ll 7--8\ theorem invertible_inverse_invertible [intro, simp]: "\ invertible u; u \ M \ \ invertible (inverse u)" using invertible_left_inverse invertible_right_inverse by blast text \p 31, l 8\ theorem invertible_inverse_inverse [simp]: "\ invertible u; u \ M \ \ inverse (inverse u) = u" by (simp add: inverse_equality) end (* monoid *) context submonoid begin text \Reasoning about @{term invertible} and @{term inverse} in submonoids.\ text \p 31, l 7\ lemma submonoid_invertible [intro, simp]: "\ sub.invertible u; u \ N \ \ invertible u" using invertibleI by blast text \p 31, l 7\ lemma submonoid_inverse_closed [intro, simp]: "\ sub.invertible u; u \ N \ \ inverse u \ N" using inverse_equality by auto end (* submonoid *) text \Def 1.2\ text \p 31, ll 9--10\ locale group = monoid G "(\)" \ for G and composition (infixl "\" 70) and unit ("\") + assumes invertible [simp, intro]: "u \ G \ invertible u" text \p 31, ll 11--12\ locale subgroup = submonoid G M "(\)" \ + sub: group G "(\)" \ for G and M and composition (infixl "\" 70) and unit ("\") begin text \Reasoning about @{term invertible} and @{term inverse} in subgroups.\ text \p 31, ll 11--12\ lemma subgroup_inverse_equality [simp]: "u \ G \ inverse u = sub.inverse u" by (simp add: inverse_equality) text \p 31, ll 11--12\ lemma subgroup_inverse_iff [simp]: "\ invertible x; x \ M \ \ inverse x \ G \ x \ G" using invertible_inverse_inverse sub.invertible_inverse_closed by fastforce end (* subgroup *) text \p 31, ll 11--12\ lemma subgroup_transitive [trans]: assumes "subgroup K H composition unit" and "subgroup H G composition unit" shows "subgroup K G composition unit" proof - interpret K: subgroup K H composition unit by fact interpret H: subgroup H G composition unit by fact show ?thesis by unfold_locales auto qed context monoid begin text \Jacobson states both directions, but the other one is trivial.\ text \p 31, ll 12--15\ theorem subgroupI: fixes G assumes subset [THEN subsetD, intro]: "G \ M" and [intro]: "\ \ G" and [intro]: "\g h. \ g \ G; h \ G \ \ g \ h \ G" and [intro]: "\g. g \ G \ invertible g" and [intro]: "\g. g \ G \ inverse g \ G" shows "subgroup G M (\) \" proof - interpret sub: monoid G "(\)" \ by unfold_locales auto show ?thesis proof unfold_locales fix u assume [intro]: "u \ G" show "sub.invertible u" using invertible_left_inverse invertible_right_inverse by blast qed auto qed text \p 31, l 16\ definition "Units = {u \ M. invertible u}" text \p 31, l 16\ lemma mem_UnitsI: "\ invertible u; u \ M \ \ u \ Units" unfolding Units_def by clarify text \p 31, l 16\ lemma mem_UnitsD: "\ u \ Units \ \ invertible u \ u \ M" unfolding Units_def by clarify text \p 31, ll 16--21\ interpretation units: subgroup Units M proof (rule subgroupI) fix u1 u2 assume Units [THEN mem_UnitsD, simp]: "u1 \ Units" "u2 \ Units" have "(u1 \ u2) \ (inverse u2 \ inverse u1) = (u1 \ (u2 \ inverse u2)) \ inverse u1" by (simp add: associative del: invertible_left_inverse invertible_right_inverse) also have "\ = \" by simp finally have inv1: "(u1 \ u2) \ (inverse u2 \ inverse u1) = \" by simp \ \ll 16--18\ have "(inverse u2 \ inverse u1) \ (u1 \ u2) = (inverse u2 \ (inverse u1 \ u1)) \ u2" by (simp add: associative del: invertible_left_inverse invertible_right_inverse) also have "\ = \" by simp - finally have inv2: "(inverse u2 \ inverse u1) \ (u1 \ u2) = \" by simp \ \l 9, “and similarly”\ + finally have inv2: "(inverse u2 \ inverse u1) \ (u1 \ u2) = \" by simp \ \l 9, ``and similarly''\ show "u1 \ u2 \ Units" using inv1 inv2 invertibleI mem_UnitsI by auto qed (auto simp: Units_def) text \p 31, ll 21--22\ theorem group_of_Units [intro, simp]: "group Units (\) \" .. text \p 31, l 19\ lemma composition_invertible [simp, intro]: "\ invertible x; invertible y; x \ M; y \ M \ \ invertible (x \ y)" using mem_UnitsD mem_UnitsI by blast text \p 31, l 20\ lemma unit_invertible: "invertible \" by fast text \Useful simplification rules\ text \p 31, l 22\ lemma invertible_right_inverse2: "\ invertible u; u \ M; v \ M \ \ u \ (inverse u \ v) = v" by (simp add: associative [THEN sym]) text \p 31, l 22\ lemma invertible_left_inverse2: "\ invertible u; u \ M; v \ M \ \ inverse u \ (u \ v) = v" by (simp add: associative [THEN sym]) text \p 31, l 22\ lemma inverse_composition_commute: assumes [simp]: "invertible x" "invertible y" "x \ M" "y \ M" shows "inverse (x \ y) = inverse y \ inverse x" proof - have "inverse (x \ y) \ (x \ y) = (inverse y \ inverse x) \ (x \ y)" by (simp add: invertible_left_inverse2 associative) then show ?thesis by (simp del: invertible_left_inverse) qed end (* monoid *) text \p 31, l 24\ context transformations begin text \p 31, ll 25--26\ theorem invertible_is_bijective: assumes dom: "\ \ S \\<^sub>E S" shows "invertible \ \ bij_betw \ S S" proof - from dom interpret map \ S S by unfold_locales show ?thesis by (auto simp add: bij_betw_iff_has_inverse invertible_def) qed text \p 31, ll 26--27\ theorem Units_bijective: "Units = {\ \ S \\<^sub>E S. bij_betw \ S S}" unfolding Units_def by (auto simp add: invertible_is_bijective) text \p 31, ll 26--27\ lemma Units_bij_betwI [intro, simp]: "\ \ Units \ bij_betw \ S S" by (simp add: Units_bijective) text \p 31, ll 26--27\ lemma Units_bij_betwD [dest, simp]: "\ \ \ S \\<^sub>E S; bij_betw \ S S \ \ \ \ Units" unfolding Units_bijective by simp text \p 31, ll 28--29\ abbreviation "Sym \ Units" text \p 31, ll 26--28\ sublocale symmetric: group "Sym" "compose S" "identity S" by (fact group_of_Units) end (* transformations *) text \p 32, ll 18--19\ locale transformation_group = transformations S + symmetric: subgroup G Sym "compose S" "identity S" for G and S begin text \p 32, ll 18--19\ lemma transformation_group_closed [intro, simp]: "\ \ \ G; x \ S \ \ \ x \ S" using bij_betwE by blast text \p 32, ll 18--19\ lemma transformation_group_undefined [intro, simp]: "\ \ \ G; x \ S \ \ \ x = undefined" by (metis compose_def symmetric.sub.right_unit restrict_apply) end (* transformation_group *) subsection \Isomorphisms. Cayley's Theorem\ text \Def 1.3\ text \p 37, ll 7--11\ locale monoid_isomorphism = bijective_map \ M M' + source: monoid M "(\)" \ + target: monoid M' "(\')" "\'" for \ and M and composition (infixl "\" 70) and unit ("\") and M' and composition' (infixl "\''" 70) and unit' ("\''") + assumes commutes_with_composition: "\ x \ M; y \ M \ \ \ x \' \ y = \ (x \ y)" and commutes_with_unit: "\ \ = \'" text \p 37, l 10\ definition isomorphic_as_monoids (infixl "\\<^sub>M" 50) where "\ \\<^sub>M \' \ (let (M, composition, unit) = \; (M', composition', unit') = \' in (\\. monoid_isomorphism \ M composition unit M' composition' unit'))" text \p 37, ll 11--12\ locale monoid_isomorphism' = bijective_map \ M M' + source: monoid M "(\)" \ + target: monoid M' "(\')" "\'" for \ and M and composition (infixl "\" 70) and unit ("\") and M' and composition' (infixl "\''" 70) and unit' ("\''") + assumes commutes_with_composition: "\ x \ M; y \ M \ \ \ x \' \ y = \ (x \ y)" text \p 37, ll 11--12\ sublocale monoid_isomorphism \ monoid_isomorphism' by unfold_locales (simp add: commutes_with_composition) text \Both definitions are equivalent.\ text \p 37, ll 12--15\ sublocale monoid_isomorphism' \ monoid_isomorphism proof unfold_locales { fix y assume "y \ M'" then obtain x where "\ x = y" "x \ M" by (metis image_iff surjective) then have "y \' \ \ = y" using commutes_with_composition by auto } then show "\ \ = \'" by fastforce qed (simp add: commutes_with_composition) context monoid_isomorphism begin text \p 37, ll 30--33\ theorem inverse_monoid_isomorphism: "monoid_isomorphism (restrict (inv_into M \) M') M' (\') \' M (\) \" using commutes_with_composition commutes_with_unit surjective by unfold_locales auto end (* monoid_isomorphism *) text \We only need that @{term \} is symmetric.\ text \p 37, ll 28--29\ theorem isomorphic_as_monoids_symmetric: "(M, composition, unit) \\<^sub>M (M', composition', unit') \ (M', composition', unit') \\<^sub>M (M, composition, unit)" by (simp add: isomorphic_as_monoids_def) (meson monoid_isomorphism.inverse_monoid_isomorphism) text \p 38, l 4\ locale left_translations_of_monoid = monoid begin (* We take the liberty of omitting "left_" from the name of the translation operation. The derived transformation monoid and group won't be qualified with "left" either. This avoids qualifications such as "left.left_...". In contexts where left and right translations are used simultaneously, notably subgroup_of_group, qualifiers are needed. *) text \p 38, ll 5--7\ definition translation ("'(_')\<^sub>L") where "translation = (\a \ M. \x \ M. a \ x)" text \p 38, ll 5--7\ lemma translation_map [intro, simp]: "a \ M \ (a)\<^sub>L \ M \\<^sub>E M" unfolding translation_def by simp text \p 38, ll 5--7\ lemma Translations_maps [intro, simp]: "translation ` M \ M \\<^sub>E M" by (simp add: image_subsetI) text \p 38, ll 5--7\ lemma translation_apply: "\ a \ M; b \ M \ \ (a)\<^sub>L b = a \ b" unfolding translation_def by auto text \p 38, ll 5--7\ lemma translation_exist: "f \ translation ` M \ \a \ M. f = (a)\<^sub>L" by auto text \p 38, ll 5--7\ lemmas Translations_E [elim] = translation_exist [THEN bexE] text \p 38, l 10\ theorem translation_unit_eq [simp]: "identity M = (\)\<^sub>L" unfolding translation_def by auto text \p 38, ll 10--11\ theorem translation_composition_eq [simp]: assumes [simp]: "a \ M" "b \ M" shows "compose M (a)\<^sub>L (b)\<^sub>L = (a \ b)\<^sub>L" unfolding translation_def by rule (simp add: associative compose_def) (* Activate @{locale monoid} to simplify subsequent proof. *) text \p 38, ll 7--9\ sublocale transformation: transformations M . text \p 38, ll 7--9\ theorem Translations_transformation_monoid: "transformation_monoid (translation ` M) M" by unfold_locales auto text \p 38, ll 7--9\ sublocale transformation: transformation_monoid "translation ` M" M by (fact Translations_transformation_monoid) text \p 38, l 12\ sublocale map translation M "translation ` M" by unfold_locales (simp add: translation_def) text \p 38, ll 12--16\ theorem translation_isomorphism [intro]: "monoid_isomorphism translation M (\) \ (translation ` M) (compose M) (identity M)" proof unfold_locales have "inj_on translation M" proof (rule inj_onI) fix a b assume [simp]: "a \ M" "b \ M" "(a)\<^sub>L = (b)\<^sub>L" have "(a)\<^sub>L \ = (b)\<^sub>L \" by simp then show "a = b" by (simp add: translation_def) qed then show "bij_betw translation M (translation ` M)" by (simp add: inj_on_imp_bij_betw) qed simp_all text \p 38, ll 12--16\ sublocale monoid_isomorphism translation M "(\)" \ "translation ` M" "compose M" "identity M" .. end (* left_translations_of_monoid *) context monoid begin text \p 38, ll 1--2\ interpretation left_translations_of_monoid .. text \p 38, ll 1--2\ theorem cayley_monoid: "\M' composition' unit'. transformation_monoid M' M \ (M, (\), \) \\<^sub>M (M', composition', unit')" by (simp add: isomorphic_as_monoids_def) (fast intro: Translations_transformation_monoid) end (* monoid *) text \p 38, l 17\ locale left_translations_of_group = group begin text \p 38, ll 17--18\ sublocale left_translations_of_monoid where M = G .. text \p 38, ll 17--18\ notation translation ("'(_')\<^sub>L") text \ The group of left translations is a subgroup of the symmetric group, hence @{term transformation.sub.invertible}. \ text \p 38, ll 20--22\ theorem translation_invertible [intro, simp]: assumes [simp]: "a \ G" shows "transformation.sub.invertible (a)\<^sub>L" proof show "compose G (a)\<^sub>L (inverse a)\<^sub>L = identity G" by simp next show "compose G (inverse a)\<^sub>L (a)\<^sub>L = identity G" by simp qed auto text \p 38, ll 19--20\ theorem translation_bijective [intro, simp]: "a \ G \ bij_betw (a)\<^sub>L G G" by (blast intro: transformation.invertible_is_bijective [THEN iffD1]) text \p 38, ll 18--20\ theorem Translations_transformation_group: "transformation_group (translation ` G) G" proof unfold_locales show "(translation ` G) \ transformation.Sym" unfolding transformation.Units_bijective by auto next fix \ assume \: "\ \ translation ` G" then obtain a where a: "a \ G" and eq: "\ = (a)\<^sub>L" .. with translation_invertible show "transformation.sub.invertible \" by simp qed auto text \p 38, ll 18--20\ sublocale transformation: transformation_group "translation ` G" G by (fact Translations_transformation_group) end (* left_translations_of_group *) context group begin text \p 38, ll 2--3\ interpretation left_translations_of_group .. text \p 38, ll 2--3\ theorem cayley_group: "\G' composition' unit'. transformation_group G' G \ (G, (\), \) \\<^sub>M (G', composition', unit')" by (simp add: isomorphic_as_monoids_def) (fast intro: Translations_transformation_group) end (* group *) text \Exercise 3\ text \p 39, ll 9--10\ locale right_translations_of_group = group begin text \p 39, ll 9--10\ definition translation ("'(_')\<^sub>R") where "translation = (\a \ G. \x \ G. x \ a)" text \p 39, ll 9--10\ abbreviation "Translations \ translation ` G" text \The isomorphism that will be established is a map different from @{term translation}.\ text \p 39, ll 9--10\ interpretation aux: map translation G Translations by unfold_locales (simp add: translation_def) text \p 39, ll 9--10\ lemma translation_map [intro, simp]: "a \ G \ (a)\<^sub>R \ G \\<^sub>E G" unfolding translation_def by simp text \p 39, ll 9--10\ lemma Translation_maps [intro, simp]: "Translations \ G \\<^sub>E G" by (simp add: image_subsetI) text \p 39, ll 9--10\ lemma translation_apply: "\ a \ G; b \ G \ \ (a)\<^sub>R b = b \ a" unfolding translation_def by auto text \p 39, ll 9--10\ lemma translation_exist: "f \ Translations \ \a \ G. f = (a)\<^sub>R" by auto text \p 39, ll 9--10\ lemmas Translations_E [elim] = translation_exist [THEN bexE] text \p 39, ll 9--10\ lemma translation_unit_eq [simp]: "identity G = (\)\<^sub>R" unfolding translation_def by auto text \p 39, ll 10--11\ lemma translation_composition_eq [simp]: assumes [simp]: "a \ G" "b \ G" shows "compose G (a)\<^sub>R (b)\<^sub>R = (b \ a)\<^sub>R" unfolding translation_def by rule (simp add: associative compose_def) text \p 39, ll 10--11\ sublocale transformation: transformations G . text \p 39, ll 10--11\ lemma Translations_transformation_monoid: "transformation_monoid Translations G" by unfold_locales auto text \p 39, ll 10--11\ sublocale transformation: transformation_monoid Translations G by (fact Translations_transformation_monoid) text \p 39, ll 10--11\ lemma translation_invertible [intro, simp]: assumes [simp]: "a \ G" shows "transformation.sub.invertible (a)\<^sub>R" proof show "compose G (a)\<^sub>R (inverse a)\<^sub>R = identity G" by simp next show "compose G (inverse a)\<^sub>R (a)\<^sub>R = identity G" by simp qed auto text \p 39, ll 10--11\ lemma translation_bijective [intro, simp]: "a \ G \ bij_betw (a)\<^sub>R G G" by (blast intro: transformation.invertible_is_bijective [THEN iffD1]) text \p 39, ll 10--11\ theorem Translations_transformation_group: "transformation_group Translations G" proof unfold_locales show "Translations \ transformation.Sym" unfolding transformation.Units_bijective by auto next fix \ assume \: "\ \ Translations" then obtain a where a: "a \ G" and eq: "\ = (a)\<^sub>R" .. with translation_invertible show "transformation.sub.invertible \" by simp qed auto text \p 39, ll 10--11\ sublocale transformation: transformation_group Translations G by (rule Translations_transformation_group) text \p 39, ll 10--11\ lemma translation_inverse_eq [simp]: assumes [simp]: "a \ G" shows "transformation.sub.inverse (a)\<^sub>R = (inverse a)\<^sub>R" proof (rule transformation.sub.inverse_equality) show "compose G (a)\<^sub>R (inverse a)\<^sub>R = identity G" by simp next show "compose G (inverse a)\<^sub>R (a)\<^sub>R = identity G" by simp qed auto text \p 39, ll 10--11\ theorem translation_inverse_monoid_isomorphism [intro]: "monoid_isomorphism (\a\G. transformation.symmetric.inverse (a)\<^sub>R) G (\) \ Translations (compose G) (identity G)" (is "monoid_isomorphism ?inv _ _ _ _ _ _") proof unfold_locales show "?inv \ G \\<^sub>E Translations" by (simp del: translation_unit_eq) next note bij_betw_compose [trans] have "bij_betw inverse G G" by (rule bij_betwI [where g = inverse]) auto also have "bij_betw translation G Translations" by (rule bij_betwI [where g = "\\\Translations. \ \"]) (auto simp: translation_apply) finally show "bij_betw ?inv G Translations" by (simp cong: bij_betw_cong add: compose_eq del: translation_unit_eq) next fix x and y assume [simp]: "x \ G" "y \ G" show "compose G (?inv x) (?inv y) = (?inv (x \ y))" by (simp add: inverse_composition_commute del: translation_unit_eq) next show "?inv \ = identity G" by (simp del: translation_unit_eq) simp qed text \p 39, ll 10--11\ sublocale monoid_isomorphism "\a\G. transformation.symmetric.inverse (a)\<^sub>R" G "(\)" \ Translations "compose G" "identity G" .. end (* right_translations_of_group *) subsection \Generalized Associativity. Commutativity\ text \p 40, l 27; p 41, ll 1--2\ locale commutative_monoid = monoid + assumes commutative: "\ x \ M; y \ M \ \ x \ y = y \ x" text \p 41, l 2\ locale abelian_group = group + commutative_monoid G "(\)" \ subsection \Orbits. Cosets of a Subgroup\ context transformation_group begin text \p 51, ll 18--20\ definition Orbit_Relation where "Orbit_Relation = {(x, y). x \ S \ y \ S \ (\\ \ G. y = \ x)}" text \p 51, ll 18--20\ lemma Orbit_Relation_memI [intro]: "\ \\ \ G. y = \ x; x \ S \ \ (x, y) \ Orbit_Relation" unfolding Orbit_Relation_def by auto text \p 51, ll 18--20\ lemma Orbit_Relation_memE [elim]: "\ (x, y) \ Orbit_Relation; \\. \ \ \ G; x \ S; y = \ x \ \ Q \ \ Q" unfolding Orbit_Relation_def by auto text \p 51, ll 20--23, 26--27\ sublocale orbit: equivalence S Orbit_Relation proof (unfold_locales, auto simp: Orbit_Relation_def) fix x assume x: "x \ S" then have id: "x = identity S x" by simp with x show "\\ \ G. x = \ x" by fast fix \ assume \: "\ \ G" with x id have y: "x = compose S (symmetric.inverse \) \ x" by auto with x \ show "\\' \ G. x = \' (\ x)" by (metis compose_eq symmetric.sub.invertible symmetric.submonoid_inverse_closed) fix \ assume \: "\ \ G" with x have "\ (\ x) = compose S \ \ x" by (simp add: compose_eq) with \ \ show "\\ \ G. \ (\ x) = \ x" by fast qed text \p 51, ll 23--24\ theorem orbit_equality: "x \ S \ orbit.Class x = {\ x | \. \ \ G}" by (simp add: orbit.Class_def) (blast intro: orbit.symmetric dest: orbit.symmetric) end (* transformation_group *) context monoid_isomorphism begin text \p 52, ll 16--17\ theorem image_subgroup: assumes "subgroup G M (\) \" shows "subgroup (\ ` G) M' (\') \'" proof - interpret subgroup G M "(\)" \ by fact interpret image: monoid "\ ` G" "(\')" "\'" by unfold_locales (auto simp add: commutes_with_composition commutes_with_unit [symmetric]) show ?thesis proof (unfold_locales, auto) fix x assume x: "x \ G" show "image.invertible (\ x)" proof show "\ (sub.inverse x) \ \ ` G" using x by simp qed (auto simp: x commutes_with_composition commutes_with_unit) qed qed end (* monoid_isomorphism *) text \ Technical device to achieve Jacobson's notation for @{text Right_Coset} and @{text Left_Coset}. The definitions are pulled out of @{text subgroup_of_group} to a context where @{text H} is not a parameter. \ text \p 52, l 20\ locale coset_notation = fixes composition (infixl "\" 70) begin text \Equation 23\ text \p 52, l 20\ definition Right_Coset (infixl "|\" 70) where "H |\ x = {h \ x | h. h \ H}" text \p 53, ll 8--9\ definition Left_Coset (infixl "\|" 70) where "x \| H = {x \ h | h. h \ H}" text \p 52, l 20\ lemma Right_Coset_memI [intro]: "h \ H \ h \ x \ H |\ x" unfolding Right_Coset_def by blast text \p 52, l 20\ lemma Right_Coset_memE [elim]: "\ a \ H |\ x; \h. \ h \ H; a = h \ x \ \ P \ \ P" unfolding Right_Coset_def by blast text \p 53, ll 8--9\ lemma Left_Coset_memI [intro]: "h \ H \ x \ h \ x \| H" unfolding Left_Coset_def by blast text \p 53, ll 8--9\ lemma Left_Coset_memE [elim]: "\ a \ x \| H; \h. \ h \ H; a = x \ h \ \ P \ \ P" unfolding Left_Coset_def by blast end (* coset_notation *) text \p 52, l 12\ locale subgroup_of_group = subgroup H G "(\)" \ + coset_notation "(\)" + group G "(\)" \ for H and G and composition (infixl "\" 70) and unit ("\") begin text \p 52, ll 12--14\ interpretation left: left_translations_of_group .. interpretation right: right_translations_of_group .. text \ @{term "left.translation ` H"} denotes Jacobson's @{text "H\<^sub>L(G)"} and @{term "left.translation ` G"} denotes Jacobson's @{text "G\<^sub>L"}. \ text \p 52, ll 16--18\ theorem left_translations_of_subgroup_are_transformation_group [intro]: "transformation_group (left.translation ` H) G" proof - have "subgroup (left.translation ` H) (left.translation ` G) (compose G) (identity G)" by (rule left.image_subgroup) unfold_locales also have "subgroup (left.translation ` G) left.transformation.Sym (compose G) (identity G)" .. finally interpret right_coset: subgroup "left.translation ` H" left.transformation.Sym "compose G" "identity G" . show ?thesis .. qed text \p 52, l 18\ interpretation transformation_group "left.translation ` H" G .. text \p 52, ll 19--20\ theorem Right_Coset_is_orbit: "x \ G \ H |\ x = orbit.Class x" using left.translation_apply by (auto simp: orbit_equality Right_Coset_def) (metis imageI sub) text \p 52, ll 24--25\ theorem Right_Coset_Union: "(\x\G. H |\ x) = G" by (simp add: Right_Coset_is_orbit) text \p 52, l 26\ theorem Right_Coset_bij: assumes G [simp]: "x \ G" "y \ G" shows "bij_betw (inverse x \ y)\<^sub>R (H |\ x) (H |\ y)" proof (rule bij_betw_imageI) show "inj_on (inverse x \ y)\<^sub>R (H |\ x)" by (fastforce intro: inj_onI simp add: Right_Coset_is_orbit right.translation_apply orbit.block_closed) next show "(inverse x \ y)\<^sub>R ` (H |\ x) = H |\ y" by (force simp add: right.translation_apply associative invertible_right_inverse2) qed text \p 52, ll 25--26\ theorem Right_Cosets_cardinality: "\ x \ G; y \ G \ \ card (H |\ x) = card (H |\ y)" by (fast intro: bij_betw_same_card Right_Coset_bij) text \p 52, l 27\ theorem Right_Coset_unit: "H |\ \ = H" by (force simp add: Right_Coset_def) text \p 52, l 27\ theorem Right_Coset_cardinality: "x \ G \ card (H |\ x) = card H" using Right_Coset_unit Right_Cosets_cardinality unit_closed by presburger text \p 52, ll 31--32\ definition "index = card orbit.Partition" text \Theorem 1.5\ text \p 52, ll 33--35; p 53, ll 1--2\ theorem lagrange: "finite G \ card G = card H * index" unfolding index_def apply (subst card_partition) apply (auto simp: finite_UnionD orbit.complete orbit.disjoint) apply (metis Right_Coset_cardinality Right_Coset_is_orbit orbit.Block_self orbit.element_exists) done end (* subgroup_of_group *) text \Left cosets\ context subgroup begin text \p 31, ll 11--12\ lemma image_of_inverse [intro, simp]: "x \ G \ x \ inverse ` G" by (metis image_eqI sub.invertible sub.invertible_inverse_closed sub.invertible_inverse_inverse subgroup_inverse_equality) end (* subgroup *) context group begin (* Does Jacobson show this somewhere? *) text \p 53, ll 6--7\ lemma inverse_subgroupI: assumes sub: "subgroup H G (\) \" shows "subgroup (inverse ` H) G (\) \" proof - from sub interpret subgroup H G "(\)" \ . interpret inv: monoid "inverse ` H" "(\)" \ by unfold_locales (auto simp del: subgroup_inverse_equality) interpret inv: group "inverse ` H" "(\)" \ by unfold_locales (force simp del: subgroup_inverse_equality) show ?thesis by unfold_locales (auto simp del: subgroup_inverse_equality) qed text \p 53, ll 6--7\ lemma inverse_subgroupD: assumes sub: "subgroup (inverse ` H) G (\) \" and inv: "H \ Units" shows "subgroup H G (\) \" proof - from sub have "subgroup (inverse ` inverse ` H) G (\) \" by (rule inverse_subgroupI) moreover from inv [THEN subsetD, simplified Units_def] have "inverse ` inverse ` H = H" by (simp cong: image_cong add: image_comp) ultimately show ?thesis by simp qed end (* group *) context subgroup_of_group begin text \p 53, l 6\ interpretation right_translations_of_group .. text \ @{term "translation ` H"} denotes Jacobson's @{text "H\<^sub>R(G)"} and @{term "Translations"} denotes Jacobson's @{text "G\<^sub>R"}. \ text \p 53, ll 6--7\ theorem right_translations_of_subgroup_are_transformation_group [intro]: "transformation_group (translation ` H) G" proof - have "subgroup ((\a\G. transformation.symmetric.inverse (a)\<^sub>R) ` H) Translations (compose G) (identity G)" by (rule image_subgroup) unfold_locales also have "subgroup Translations transformation.Sym (compose G) (identity G)" .. finally interpret left_coset: subgroup "translation ` H" transformation.Sym "compose G" "identity G" by (auto intro: transformation.symmetric.inverse_subgroupD cong: image_cong simp: image_image transformation.symmetric.Units_def simp del: translation_unit_eq) show ?thesis .. qed text \p 53, ll 6--7\ interpretation transformation_group "translation ` H" G .. text \Equation 23 for left cosets\ text \p 53, ll 7--8\ theorem Left_Coset_is_orbit: "x \ G \ x \| H = orbit.Class x" using translation_apply by (auto simp: orbit_equality Left_Coset_def) (metis imageI sub) end (* subgroup_of_group *) subsection \Congruences. Quotient Monoids and Groups\ text \Def 1.4\ text \p 54, ll 19--22\ locale monoid_congruence = monoid + equivalence where S = M + assumes cong: "\ (a, a') \ E; (b, b') \ E \ \ (a \ b, a' \ b') \ E" begin text \p 54, ll 26--28\ theorem Class_cong: "\ Class a = Class a'; Class b = Class b'; a \ M; a' \ M; b \ M; b' \ M \ \ Class (a \ b) = Class (a' \ b')" by (simp add: Class_equivalence cong) text \p 54, ll 28--30\ definition quotient_composition (infixl "[\]" 70) where "quotient_composition = (\A \ M / E. \B \ M / E. THE C. \a \ A. \b \ B. C = Class (a \ b))" text \p 54, ll 28--30\ theorem Class_commutes_with_composition: "\ a \ M; b \ M \ \ Class a [\] Class b = Class (a \ b)" by (auto simp: quotient_composition_def intro: Class_cong [OF Class_eq Class_eq] del: equalityI) text \p 54, ll 30--31\ theorem quotient_composition_closed [intro, simp]: "\ A \ M / E; B \ M / E \ \ A [\] B \ M / E" by (erule quotient_ClassE)+ (simp add: Class_commutes_with_composition) text \p 54, l 32; p 55, ll 1--3\ sublocale quotient: monoid "M / E" "([\])" "Class \" by unfold_locales (auto simp: Class_commutes_with_composition associative elim!: quotient_ClassE) end (* monoid_congruence *) text \p 55, ll 16--17\ locale group_congruence = group + monoid_congruence where M = G begin text \p 55, ll 16--17\ notation quotient_composition (infixl "[\]" 70) text \p 55, l 18\ theorem Class_right_inverse: "a \ G \ Class a [\] Class (inverse a) = Class \" by (simp add: Class_commutes_with_composition) text \p 55, l 18\ theorem Class_left_inverse: "a \ G \ Class (inverse a) [\] Class a = Class \" by (simp add: Class_commutes_with_composition) text \p 55, l 18\ theorem Class_invertible: "a \ G \ quotient.invertible (Class a)" by (blast intro!: Class_right_inverse Class_left_inverse)+ text \p 55, l 18\ theorem Class_commutes_with_inverse: "a \ G \ quotient.inverse (Class a) = Class (inverse a)" by (rule quotient.inverse_equality) (auto simp: Class_right_inverse Class_left_inverse) text \p 55, l 17\ sublocale quotient: group "G / E" "([\])" "Class \" by unfold_locales (metis Block_self Class_invertible element_exists) end (* group_congruence *) text \Def 1.5\ text \p 55, ll 22--25\ locale normal_subgroup = subgroup_of_group K G "(\)" \ for K and G and composition (infixl "\" 70) and unit ("\") + assumes normal: "\ g \ G; k \ K \ \ inverse g \ k \ g \ K" text \Lemmas from the proof of Thm 1.6\ context subgroup_of_group begin text \We use @{term H} for @{term K}.\ text \p 56, ll 14--16\ theorem Left_equals_Right_coset_implies_normality: assumes [simp]: "\g. g \ G \ g \| H = H |\ g" shows "normal_subgroup H G (\) \" proof fix g k assume [simp]: "g \ G" "k \ H" have "k \ g \ g \| H" by auto then obtain k' where "k \ g = g \ k'" and "k' \ H" by blast then show "inverse g \ k \ g \ H" by (simp add: associative invertible_left_inverse2) qed end (* subgroup_of_group *) text \Thm 1.6, first part\ context group_congruence begin text \Jacobson's $K$\ text \p 56, l 29\ definition "Normal = Class \" text \p 56, ll 3--6\ interpretation subgroup "Normal" G "(\)" \ unfolding Normal_def proof (rule subgroupI) fix k1 and k2 assume K: "k1 \ Class \" "k2 \ Class \" then have "k1 \ k2 \ Class (k1 \ k2)" by blast also have "\ = Class k1 [\] Class k2" using K by (auto simp add: Class_commutes_with_composition Class_closed) also have "\ = Class \ [\] Class \" using K by (metis ClassD Class_eq unit_closed) also have "\ = Class \" by simp finally show "k1 \ k2 \ Class \" . next fix k assume K: "k \ Class \" then have "inverse k \ Class (inverse k)" by blast also have "\ = quotient.inverse (Class k)" using Class_commutes_with_inverse K by blast also have "\ = quotient.inverse (Class \)" using Block_self K by auto also have "\ = Class \" using quotient.inverse_unit by blast finally show "inverse k \ Class \" . qed auto text \Coset notation\ text \p 56, ll 5--6\ interpretation subgroup_of_group "Normal" G "(\)" \ .. text \Equation 25 for right cosets\ text \p 55, ll 29--30; p 56, ll 6--11\ theorem Right_Coset_Class_unit: assumes g: "g \ G" shows "Normal |\ g = Class g" unfolding Normal_def proof auto fix a \ \ll 6--8\ assume a: "a \ Class g" from a g have "a \ inverse g \ Class (a \ inverse g)" by blast also from a g have "\ = Class a [\] Class (inverse g)" by (simp add: Class_commutes_with_composition block_closed) also from a g have "\ = Class g [\] quotient.inverse (Class g)" using Block_self Class_commutes_with_inverse by auto also from g have "\ = Class \" by simp finally show "a \ Class \ |\ g" unfolding Right_Coset_def by simp (metis Class_closed a associative g inverse_equality invertible invertible_def right_unit) next fix a \ \ll 8--9\ assume a: "a \ Class \ |\ g" then obtain k where eq: "a = k \ g" and k: "k \ Class \" by blast with g have "Class a = Class k [\] Class g" using Class_commutes_with_composition by auto also from k have "\ = Class \ [\] Class g" using Block_self by auto also from g have "\ = Class g" by simp finally show "a \ Class g" using g eq k composition_closed quotient.unit_closed by blast qed text \Equation 25 for left cosets\ text \p 55, ll 29--30; p 56, ll 6--11\ theorem Left_Coset_Class_unit: assumes g: "g \ G" shows "g \| Normal = Class g" unfolding Normal_def proof auto fix a \ \ll 6--8\ assume a: "a \ Class g" from a g have "inverse g \ a \ Class (inverse g \ a)" by blast also from a g have "\ = Class (inverse g) [\] Class a" by (simp add: Class_commutes_with_composition block_closed) also from a g have "\ = quotient.inverse (Class g) [\] Class g" using Block_self Class_commutes_with_inverse by auto also from g have "\ = Class \" by simp finally show "a \ g \| Class \" unfolding Left_Coset_def by simp (metis Class_closed a associative g inverse_equality invertible invertible_def right_unit) next fix a \ \ll 8--9, ``the same thing holds''\ assume a: "a \ g \| Class \" then obtain k where eq: "a = g \ k" and k: "k \ Class \" by blast with g have "Class a = Class g [\] Class k" using Class_commutes_with_composition by auto also from k have "\ = Class g [\] Class \" using Block_self by auto also from g have "\ = Class g" by simp finally show "a \ Class g" using g eq k composition_closed quotient.unit_closed by blast qed text \Thm 1.6, statement of first part\ text \p 55, ll 28--29; p 56, ll 12--16\ theorem Class_unit_is_normal: "normal_subgroup Normal G (\) \" proof - { fix g assume "g \ G" then have "g \| Normal = Normal |\ g" by (simp add: Right_Coset_Class_unit Left_Coset_Class_unit) } then show ?thesis by (rule Left_equals_Right_coset_implies_normality) qed sublocale normal: normal_subgroup Normal G "(\)" \ by (fact Class_unit_is_normal) end (* group_congruence *) context normal_subgroup begin text \p 56, ll 16--19\ theorem Left_equals_Right_coset: "g \ G \ g \| K = K |\ g" proof assume [simp]: "g \ G" show "K |\ g \ g \| K" proof fix x assume x: "x \ K |\ g" then obtain k where "x = k \ g" and [simp]: "k \ K" by (auto simp add: Right_Coset_def) then have "x = g \ (inverse g \ k \ g)" by (simp add: associative invertible_right_inverse2) also from normal have "\ \ g \| K" by (auto simp add: Left_Coset_def) finally show "x \ g \| K" . qed next assume [simp]: "g \ G" show "g \| K \ K |\ g" proof fix x assume x: "x \ g \| K" then obtain k where "x = g \ k" and [simp]: "k \ K" by (auto simp add: Left_Coset_def) then have "x = (inverse (inverse g) \ k \ inverse g) \ g" by (simp add: associative del: invertible_right_inverse) also from normal [where g = "inverse g"] have "\ \ K |\ g" by (auto simp add: Right_Coset_def) finally show "x \ K |\ g" . qed qed text \Thm 1.6, second part\ text \p 55, ll 31--32; p 56, ll 20--21\ definition "Congruence = {(a, b). a \ G \ b \ G \ inverse a \ b \ K}" text \p 56, ll 21--22\ interpretation right_translations_of_group .. text \p 56, ll 21--22\ interpretation transformation_group "translation ` K" G rewrites "Orbit_Relation = Congruence" proof - interpret transformation_group "translation ` K" G .. show "Orbit_Relation = Congruence" unfolding Orbit_Relation_def Congruence_def by (force simp: invertible_left_inverse2 invertible_right_inverse2 translation_apply simp del: restrict_apply) qed rule text \p 56, ll 20--21\ lemma CongruenceI: "\ a = b \ k; a \ G; b \ G; k \ K \ \ (a, b) \ Congruence" by (clarsimp simp: Congruence_def associative inverse_composition_commute) text \p 56, ll 20--21\ lemma CongruenceD: "(a, b) \ Congruence \ \k\K. a = b \ k" by (drule orbit.symmetric) (force simp: Congruence_def invertible_right_inverse2) text \ ``We showed in the last section that the relation we are considering is an equivalence relation in @{term G} for any subgroup @{term K} of @{term G}. We now proceed to show that normality of @{term K} ensures that [...] $a \equiv b \pmod{K}$ is a congruence.'' \ text \p 55, ll 30--32; p 56, ll 1, 22--28\ sublocale group_congruence where E = Congruence rewrites "Normal = K" proof - show "group_congruence G (\) \ Congruence" proof unfold_locales note CongruenceI [intro] CongruenceD [dest] fix a g b h assume 1: "(a, g) \ Congruence" and 2: "(b, h) \ Congruence" then have G: "a \ G" "g \ G" "b \ G" "h \ G" unfolding Congruence_def by clarify+ from 1 obtain k1 where a: "a = g \ k1" and k1: "k1 \ K" by blast from 2 obtain k2 where b: "b = h \ k2" and k2: "k2 \ K" by blast from G Left_equals_Right_coset have "K |\ h = h \| K" by blast with k1 obtain k3 where c: "k1 \ h = h \ k3" and k3: "k3 \ K" unfolding Left_Coset_def Right_Coset_def by blast from G k1 k2 a b have "a \ b = g \ k1 \ h \ k2" by (simp add: associative) also from G k1 k3 c have "\ = g \ h \ k3 \ k2" by (simp add: associative) also have "\ = (g \ h) \ (k3 \ k2)" using G k2 k3 by (simp add: associative) finally show "(a \ b, g \ h) \ Congruence" using G k2 k3 by blast qed then interpret group_congruence where E = Congruence . show "Normal = K" unfolding Normal_def orbit.Class_def unfolding Congruence_def using invertible_inverse_inverse submonoid_inverse_closed by fastforce qed end (* normal_subgroup *) (* deletes translations and orbits, recovers Class for congruence class *) context group begin text \Pulled out of @{locale normal_subgroup} to achieve standard notation.\ text \p 56, ll 31--32\ abbreviation Factor_Group (infixl "'/'/" 75) where "S // K \ S / (normal_subgroup.Congruence K G (\) \)" end (* group *) context normal_subgroup begin text \p 56, ll 28--29\ theorem Class_unit_normal_subgroup: "Class \ = K" unfolding Class_def unfolding Congruence_def using invertible_inverse_inverse submonoid_inverse_closed by fastforce text \p 56, ll 1--2; p 56, l 29\ theorem Class_is_Left_Coset: "g \ G \ Class g = g \| K" using Left_Coset_Class_unit Class_unit_normal_subgroup by simp text \p 56, l 29\ lemma Left_CosetE: "\ A \ G // K; \a. a \ G \ P (a \| K) \ \ P A" by (metis Class_is_Left_Coset quotient_ClassE) text \Equation 26\ text \p 56, ll 32--34\ theorem factor_composition [simp]: "\ g \ G; h \ G \ \ (g \| K) [\] (h \| K) = g \ h \| K" using Class_commutes_with_composition Class_is_Left_Coset by auto text \p 56, l 35\ theorem factor_unit: "K = \ \| K" using Class_is_Left_Coset Class_unit_normal_subgroup by blast text \p 56, l 35\ theorem factor_inverse [simp]: "g \ G \ quotient.inverse (g \| K) = (inverse g \| K)" using Class_commutes_with_inverse Class_is_Left_Coset by auto end (* normal_subgroup *) text \p 57, ll 4--5\ locale subgroup_of_abelian_group = subgroup_of_group H G "(\)" \ + abelian_group G "(\)" \ for H and G and composition (infixl "\" 70) and unit ("\") text \p 57, ll 4--5\ sublocale subgroup_of_abelian_group \ normal_subgroup H G "(\)" \ using commutative invertible_right_inverse2 by unfold_locales auto subsection \Homomorphims\ text \Def 1.6\ text \p 58, l 33; p 59, ll 1--2\ locale monoid_homomorphism = map \ M M'+ source: monoid M "(\)" \ + target: monoid M' "(\')" "\'" for \ and M and composition (infixl "\" 70) and unit ("\") and M' and composition' (infixl "\''" 70) and unit' ("\''") + assumes commutes_with_composition: "\ x \ M; y \ M \ \ \ (x \ y) = \ x \' \ y" and commutes_with_unit: "\ \ = \'" begin text \Jacobson notes that @{thm [source] commutes_with_unit} is not necessary for groups, but doesn't make use of that later.\ text \p 58, l 33; p 59, ll 1--2\ notation source.invertible ("invertible _" [100] 100) notation source.inverse ("inverse _" [100] 100) notation target.invertible ("invertible'' _" [100] 100) notation target.inverse ("inverse'' _" [100] 100) end (* monoid_homomorphism *) text \p 59, ll 29--30\ locale monoid_epimorphism = monoid_homomorphism + surjective_map \ M M' text \p 59, l 30\ locale monoid_monomorphism = monoid_homomorphism + injective_map \ M M' text \p 59, ll 30--31\ sublocale monoid_isomorphism \ monoid_epimorphism by unfold_locales (auto simp: commutes_with_composition commutes_with_unit) text \p 59, ll 30--31\ sublocale monoid_isomorphism \ monoid_monomorphism by unfold_locales (auto simp: commutes_with_composition commutes_with_unit) context monoid_homomorphism begin text \p 59, ll 33--34\ theorem invertible_image_lemma: assumes "invertible a" "a \ M" shows "\ a \' \ (inverse a) = \'" and "\ (inverse a) \' \ a = \'" using assms commutes_with_composition commutes_with_unit source.inverse_equality by auto (metis source.invertible_inverse_closed source.invertible_left_inverse) text \p 59, l 34; p 60, l 1\ theorem invertible_target_invertible [intro, simp]: "\ invertible a; a \ M \ \ invertible' (\ a)" using invertible_image_lemma by blast text \p 60, l 1\ theorem invertible_commutes_with_inverse: "\ invertible a; a \ M \ \ \ (inverse a) = inverse' (\ a)" using invertible_image_lemma target.inverse_equality by fastforce end (* monoid_homomorphism *) text \p 60, ll 32--34; p 61, l 1\ sublocale monoid_congruence \ natural: monoid_homomorphism Class M "(\)" \ "M / E" "([\])" "Class \" by unfold_locales (auto simp: PiE_I Class_commutes_with_composition) text \Fundamental Theorem of Homomorphisms of Monoids\ text \p 61, ll 5, 14--16\ sublocale monoid_homomorphism \ image: submonoid "\ ` M" M' "(\')" "\'" by unfold_locales (auto simp: commutes_with_composition [symmetric] commutes_with_unit [symmetric]) text \p 61, l 4\ locale monoid_homomorphism_fundamental = monoid_homomorphism begin text \p 61, ll 17--18\ sublocale fiber_relation \ M M' .. notation Fiber_Relation ("E'(_')") text \p 61, ll 6--7, 18--20\ sublocale monoid_congruence where E = "E(\)" using Class_eq by unfold_locales (rule Class_equivalence [THEN iffD1], auto simp: left_closed right_closed commutes_with_composition Fiber_equality) text \p 61, ll 7--9\ text \ @{term induced} denotes Jacobson's $\bar{\eta}$. We have the commutativity of the diagram, where @{term induced} is unique: @{thm [display] factorization} @{thm [display] uniqueness}. \ text \p 61, l 20\ notation quotient_composition (infixl "[\]" 70) text \p 61, ll 7--8, 22--25\ sublocale induced: monoid_homomorphism induced "M / E(\)" "([\])" "Class \" "M'" "(\')" "\'" apply unfold_locales apply (auto simp: commutes_with_unit) apply (fastforce simp: commutes_with_composition commutes_with_unit Class_commutes_with_composition) done text \p 61, ll 9, 26\ sublocale natural: monoid_epimorphism Class M "(\)" \ "M / E(\)" "([\])" "Class \" .. text \p 61, ll 9, 26--27\ sublocale induced: monoid_monomorphism induced "M / E(\)" "([\])" "Class \" "M'" "(\')" "\'" .. end (* monoid_homomorphism_fundamental *) text \p 62, ll 12--13\ locale group_homomorphism = monoid_homomorphism \ G "(\)" \ G' "(\')" "\'" + source: group G "(\)" \ + target: group G' "(\')" "\'" for \ and G and composition (infixl "\" 70) and unit ("\") and G' and composition' (infixl "\''" 70) and unit' ("\''") begin text \p 62, l 13\ sublocale image: subgroup "\ ` G" G' "(\')" "\'" using invertible_image_lemma by unfold_locales auto text \p 62, ll 13--14\ definition "Ker = \ -` {\'} \ G" text \p 62, ll 13--14\ lemma Ker_equality: "Ker = {a | a. a \ G \ \ a = \'}" unfolding Ker_def by auto text \p 62, ll 13--14\ lemma Ker_closed [intro, simp]: "a \ Ker \ a \ G" unfolding Ker_def by simp text \p 62, ll 13--14\ lemma Ker_image [intro]: (* loops as a simprule *) "a \ Ker \ \ a = \'" unfolding Ker_def by simp text \p 62, ll 13--14\ lemma Ker_memI [intro]: (* loops as a simprule *) "\ \ a = \'; a \ G \ \ a \ Ker" unfolding Ker_def by simp text \p 62, ll 15--16\ sublocale kernel: normal_subgroup Ker G proof - interpret kernel: submonoid Ker G unfolding Ker_def by unfold_locales (auto simp: commutes_with_composition commutes_with_unit) interpret kernel: subgroup Ker G by unfold_locales (force intro: source.invertible_right_inverse simp: Ker_image invertible_commutes_with_inverse) show "normal_subgroup Ker G (\) \" apply unfold_locales unfolding Ker_def by (auto simp: commutes_with_composition invertible_image_lemma(2)) qed text \p 62, ll 17--20\ theorem injective_iff_kernel_unit: "inj_on \ G \ Ker = {\}" proof (rule Not_eq_iff [THEN iffD1, OF iffI]) assume "Ker \ {\}" then obtain b where b: "b \ Ker" "b \ \" by blast then have "\ b = \ \" by (simp add: Ker_image) with b show "\ inj_on \ G" by (meson inj_onD kernel.sub source.unit_closed) next assume "\ inj_on \ G" then obtain a b where "a \ b" and ab: "a \ G" "b \ G" "\ a = \ b" by (meson inj_onI) then have "inverse a \ b \ \" "\ (inverse a \ b) = \'" using ab source.invertible_right_inverse2 by force (metis ab commutes_with_composition invertible_image_lemma(2) source.invertible source.invertible_inverse_closed) then have "inverse a \ b \ Ker" using Ker_memI ab by blast then show "Ker \ {\}" using \inverse a \ b \ \\ by blast qed end (* group_homomorphism *) text \p 62, l 24\ locale group_epimorphism = group_homomorphism + monoid_epimorphism \ G "(\)" \ G' "(\')" "\'" text \p 62, l 21\ locale normal_subgroup_in_kernel = group_homomorphism + contained: normal_subgroup L G "(\)" \ for L + assumes subset: "L \ Ker" begin text \p 62, l 21\ notation contained.quotient_composition (infixl "[\]" 70) text \"homomorphism onto @{term "G // L"}"\ text \p 62, ll 23--24\ sublocale natural: group_epimorphism contained.Class G "(\)" \ "G // L" "([\])" "contained.Class \" .. text \p 62, ll 25--26\ theorem left_coset_equality: assumes eq: "a \| L = b \| L" and [simp]: "a \ G" and b: "b \ G" shows "\ a = \ b" proof - obtain l where l: "b = a \ l" "l \ L" by (metis b contained.Class_is_Left_Coset contained.Class_self eq kernel.Left_Coset_memE) then have "\ a = \ a \' \ l" using Ker_image monoid_homomorphism.commutes_with_composition subset by auto also have "\ = \ b" by (simp add: commutes_with_composition l) finally show ?thesis . qed text \$\bar{\eta}$\ text \p 62, ll 26--27\ definition "induced = (\A \ G // L. THE b. \a \ G. a \| L = A \ b = \ a)" text \p 62, ll 26--27\ lemma induced_closed [intro, simp]: assumes [simp]: "A \ G // L" shows "induced A \ G'" proof - obtain a where a: "a \ G" "a \| L = A" using contained.Class_is_Left_Coset contained.Partition_def assms by auto have "(THE b. \a \ G. a \| L = A \ b = \ a) \ G'" apply (rule theI2) using a by (auto intro: left_coset_equality) then show ?thesis unfolding induced_def by simp qed text \p 62, ll 26--27\ lemma induced_undefined [intro, simp]: "A \ G // L \ induced A = undefined" unfolding induced_def by simp text \p 62, ll 26--27\ theorem induced_left_coset_closed [intro, simp]: "a \ G \ induced (a \| L) \ G'" using contained.Class_is_Left_Coset contained.Class_in_Partition by auto text \p 62, ll 26--27\ theorem induced_left_coset_equality [simp]: assumes [simp]: "a \ G" shows "induced (a \| L) = \ a" proof - have "(THE b. \a' \ G. a' \| L = a \| L \ b = \ a') = \ a" by (rule the_equality) (auto intro: left_coset_equality) then show ?thesis unfolding induced_def using contained.Class_is_Left_Coset contained.Class_in_Partition by auto qed text \p 62, l 27\ theorem induced_Left_Coset_commutes_with_composition [simp]: "\ a \ G; b \ G \ \ induced ((a \| L) [\] (b \| L)) = induced (a \| L) \' induced (b \| L)" by (simp add: commutes_with_composition) text \p 62, ll 27--28\ theorem induced_group_homomorphism: "group_homomorphism induced (G // L) ([\]) (contained.Class \) G' (\') \'" apply unfold_locales apply (auto elim!: contained.Left_CosetE simp: commutes_with_composition commutes_with_unit) using contained.factor_unit induced_left_coset_equality apply (fastforce simp: contained.Class_unit_normal_subgroup) done text \p 62, l 28\ sublocale induced: group_homomorphism induced "G // L" "([\])" "contained.Class \" G' "(\')" "\'" by (fact induced_group_homomorphism) text \p 62, ll 28--29\ theorem factorization_lemma: "a \ G \ compose G induced contained.Class a = \ a" unfolding compose_def by (simp add: contained.Class_is_Left_Coset) text \p 62, ll 29--30\ theorem factorization [simp]: "compose G induced contained.Class = \" by rule (simp add: compose_def contained.Class_is_Left_Coset map_undefined) text \ Jacobson does not state the uniqueness of @{term induced} explicitly but he uses it later, for rings, on p 107. \ text \p 62, l 30\ theorem uniqueness: assumes map: "\ \ G // L \\<^sub>E G'" and factorization: "compose G \ contained.Class = \" shows "\ = induced" proof fix A show "\ A = induced A" proof (cases "A \ G // L") case True then obtain a where [simp]: "A = contained.Class a" "a \ G" by fast then have "\ (contained.Class a) = \ a" by (metis compose_eq factorization) also have "\ = induced (contained.Class a)" by (simp add: contained.Class_is_Left_Coset) finally show ?thesis by simp qed (simp add: induced_def PiE_arb [OF map]) qed text \p 62, l 31\ theorem induced_image: "induced ` (G // L) = \ ` G" by (metis factorization contained.natural.surjective surj_compose) text \p 62, l 33\ interpretation L: normal_subgroup L Ker by unfold_locales (auto simp: subset, metis kernel.sub kernel.subgroup_inverse_equality contained.normal) text \p 62, ll 31--33\ theorem induced_kernel: "induced.Ker = Ker / L.Congruence" (* Ker // L is apparently not the right thing *) proof - have "induced.Ker = { a \| L | a. a \ G \ a \ Ker }" unfolding induced.Ker_equality by simp (metis (opaque_lifting) contained.Class_is_Left_Coset Ker_image Ker_memI induced_left_coset_equality contained.Class_in_Partition contained.representant_exists) also have "\ = Ker / L.Congruence" using L.Class_is_Left_Coset L.Class_in_Partition by auto (metis L.Class_is_Left_Coset L.representant_exists kernel.sub) finally show ?thesis . qed text \p 62, ll 34--35\ theorem induced_inj_on: "inj_on induced (G // L) \ L = Ker" apply (simp add: induced.injective_iff_kernel_unit induced_kernel contained.Class_unit_normal_subgroup) apply rule using L.block_exists apply auto [1] using L.Block_self L.Class_unit_normal_subgroup L.quotient.unit_closed L.representant_exists apply auto done end (* normal_subgroup_in_kernel *) text \Fundamental Theorem of Homomorphisms of Groups\ text \p 63, l 1\ locale group_homomorphism_fundamental = group_homomorphism begin text \p 63, l 1\ notation kernel.quotient_composition (infixl "[\]" 70) text \p 63, l 1\ sublocale normal_subgroup_in_kernel where L = Ker by unfold_locales rule text \p 62, ll 36--37; p 63, l 1\ text \ @{term induced} denotes Jacobson's $\bar{\eta}$. We have the commutativity of the diagram, where @{term induced} is unique: @{thm [display] factorization} @{thm [display] uniqueness} \ end (* group_homomorphism_fundamental *) text \p 63, l 5\ locale group_isomorphism = group_homomorphism + bijective_map \ G G' begin text \p 63, l 5\ sublocale monoid_isomorphism \ G "(\)" \ G' "(\')" "\'" by unfold_locales (auto simp: commutes_with_composition) text \p 63, l 6\ lemma inverse_group_isomorphism: "group_isomorphism (restrict (inv_into G \) G') G' (\') \' G (\) \" using commutes_with_composition commutes_with_unit surjective by unfold_locales auto end (* group_isomorphism *) text \p 63, l 6\ definition isomorphic_as_groups (infixl "\\<^sub>G" 50) where "\ \\<^sub>G \' \ (let (G, composition, unit) = \; (G', composition', unit') = \' in (\\. group_isomorphism \ G composition unit G' composition' unit'))" text \p 63, l 6\ lemma isomorphic_as_groups_symmetric: "(G, composition, unit) \\<^sub>G (G', composition', unit') \ (G', composition', unit') \\<^sub>G (G, composition, unit)" by (simp add: isomorphic_as_groups_def) (meson group_isomorphism.inverse_group_isomorphism) text \p 63, l 1\ sublocale group_isomorphism \ group_epimorphism .. text \p 63, l 1\ locale group_epimorphism_fundamental = group_homomorphism_fundamental + group_epimorphism begin text \p 63, ll 1--2\ interpretation image: group_homomorphism induced "G // Ker" "([\])" "kernel.Class \" "(\ ` G)" "(\')" "\'" by (simp add: surjective group_homomorphism_fundamental.intro induced_group_homomorphism) text \p 63, ll 1--2\ sublocale image: group_isomorphism induced "G // Ker" "([\])" "kernel.Class \" "(\ ` G)" "(\')" "\'" using induced_group_homomorphism by unfold_locales (auto simp: bij_betw_def induced_image induced_inj_on induced.commutes_with_composition) end (* group_epimorphism_fundamental *) context group_homomorphism begin text \p 63, ll 5--7\ theorem image_isomorphic_to_factor_group: "\K composition unit. normal_subgroup K G (\) \ \ (\ ` G, (\'), \') \\<^sub>G (G // K, composition, unit)" proof - interpret image: group_epimorphism_fundamental where G' = "\ ` G" by unfold_locales (auto simp: commutes_with_composition) have "group_isomorphism image.induced (G // Ker) ([\]) (kernel.Class \) (\ ` G) (\') \'" .. then have "(\ ` G, (\'), \') \\<^sub>G (G // Ker, ([\]), kernel.Class \)" by (simp add: isomorphic_as_groups_def) (meson group_isomorphism.inverse_group_isomorphism) moreover have "normal_subgroup Ker G (\) \" .. ultimately show ?thesis by blast qed end (* group_homomorphism *) end