diff --git a/thys/Complx/OG_Soundness.thy b/thys/Complx/OG_Soundness.thy --- a/thys/Complx/OG_Soundness.thy +++ b/thys/Complx/OG_Soundness.thy @@ -1,2032 +1,2029 @@ (* * Copyright 2016, Data61, CSIRO * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(DATA61_BSD) *) section \Soundness proof of Owicki-Gries w.r.t. COMPLX small-step semantics\ theory OG_Soundness imports OG_Hoare SeqCatch_decomp begin lemma pre_weaken_pre: " x \ pre P \ x \ pre (weaken_pre P P')" by (induct P, clarsimp+) lemma oghoare_Skip[rule_format, OF _ refl]: "\, \ \\<^bsub>/F\<^esub> P c Q,A \ c = Skip \ (\P'. P = AnnExpr P' \ P' \ Q)" apply (induct rule: oghoare_induct, simp_all) apply clarsimp apply (rename_tac \ \ F P Q A P' Q' A' P'') apply(case_tac P, simp_all) by force lemma oghoare_Throw[rule_format, OF _ refl]: "\, \ \\<^bsub>/F\<^esub> P c Q,A \ c = Throw \ (\P'. P = AnnExpr P' \ P' \ A)" apply (induct rule: oghoare_induct, simp_all) apply clarsimp apply (rename_tac \ \ F P Q A P' Q' A' P'') apply(case_tac P, simp_all) by force lemma oghoare_Basic[rule_format, OF _ refl]: "\, \ \\<^bsub>/F\<^esub> P c Q,A \ c = Basic f \ (\P'. P = AnnExpr P' \ P' \ {x. f x \ Q})" apply (induct rule: oghoare_induct, simp_all) apply clarsimp apply (rename_tac \ \ F P Q A P' Q' A' P'') apply(case_tac P, simp_all) by force lemma oghoare_Spec[rule_format, OF _ refl]: "\, \ \\<^bsub>/F\<^esub> P c Q,A \ c = Spec r \ (\P'. P = AnnExpr P' \ P' \ {s. (\t. (s, t) \ r \ t \ Q) \ (\t. (s, t) \ r)})" apply (induct rule: oghoare_induct, simp_all) apply clarsimp apply (rename_tac \ \ F P Q A P' Q' A' P'') apply(case_tac P, simp_all) by force lemma oghoare_DynCom[rule_format, OF _ refl]: "\, \ \\<^bsub>/F\<^esub> P c Q,A \ c = (DynCom c') \ (\r ad. P = (AnnRec r ad) \ r \ pre ad \ (\s\r. \, \ \\<^bsub>/F\<^esub> ad (c' s) Q,A))" apply (induct rule: oghoare_induct, simp_all) apply clarsimp apply clarsimp apply (rename_tac \ \ F P Q A P' Q' A' P'' x) apply(case_tac P, simp_all) apply clarsimp apply (rename_tac P s) apply (drule_tac x=s in bspec, simp) apply (rule Conseq) apply (rule_tac x="{}" in exI) apply (fastforce) done lemma oghoare_Guard[rule_format, OF _ refl]: "\,\\\<^bsub>/F\<^esub> P c Q,A \ c = Guard f g d \ (\P' r . P = AnnRec r P' \ \,\\\<^bsub>/F\<^esub> P' d Q,A \ r \ g \ pre P' \ (r \ -g \ {} \ f \ F))" apply (induct rule: oghoare_induct, simp_all) apply force apply clarsimp apply (rename_tac \ \ F P Q A P' Q' A' P'' r) apply (case_tac P, simp_all) apply clarsimp apply (rename_tac r) apply(rule conjI) apply (rule Conseq) apply (rule_tac x="{}" in exI) apply (rule_tac x="Q'" in exI) apply (rule_tac x="A'" in exI) apply (clarsimp) apply (case_tac "(r \ P') \ g \ {}") apply fast apply clarsimp apply(drule equalityD1, force) done lemma oghoare_Await[rule_format, OF _ refl]: "\, \\\<^bsub>/F\<^esub> P x Q,A \ \b c. x = Await b c \ (\r P' Q' A'. P = AnnRec r P' \ \, \\\<^bsub>/F\<^esub>(r \ b) P' c Q',A' \ atom_com c \ Q' \ Q \ A' \ A)" apply (induct rule: oghoare_induct, simp_all) apply (rename_tac \ \ F r P Q A) apply (rule_tac x=Q in exI) apply (rule_tac x=A in exI) apply clarsimp apply (rename_tac \ \ F P c Q A) apply clarsimp apply(case_tac P, simp_all) apply (rename_tac P'' Q'' A'' x y) apply (rule_tac x=Q'' in exI) apply (rule_tac x=A'' in exI) apply clarsimp apply (rule conjI[rotated]) apply blast apply (erule SeqConseq[rotated]) apply simp apply simp apply blast done lemma oghoare_Seq[rule_format, OF _ refl]: "\, \ \\<^bsub>/F\<^esub> P x Q,A \ \p1 p2. x = Seq p1 p2 \ (\ P\<^sub>1 P\<^sub>2 P' Q' A'. P = AnnComp P\<^sub>1 P\<^sub>2 \ \, \ \\<^bsub>/F\<^esub> P\<^sub>1 p1 P', A' \ P' \ pre P\<^sub>2 \ \, \ \\<^bsub>/F\<^esub> P\<^sub>2 p2 Q',A' \ Q' \ Q \ A' \ A)" apply (induct rule: oghoare_induct, simp_all) apply blast apply (rename_tac \ \ F P c Q A) apply clarsimp apply (rename_tac P'' Q'' A'') apply(case_tac P, simp_all) apply clarsimp apply (rule_tac x="P''" in exI) apply (rule_tac x="Q''" in exI) apply (rule_tac x="A''" in exI) apply clarsimp apply (rule conjI) apply (rule Conseq) apply (rule_tac x="P'" in exI) apply (rule_tac x="P''" in exI) apply (rule_tac x="A''" in exI) apply simp apply fastforce done lemma oghoare_Catch[rule_format, OF _ refl]: "\, \ \\<^bsub>/F\<^esub> P x Q,A \ \p1 p2. x = Catch p1 p2 \ (\ P\<^sub>1 P\<^sub>2 P' Q' A'. P = AnnComp P\<^sub>1 P\<^sub>2 \ \, \ \\<^bsub>/F\<^esub> P\<^sub>1 p1 Q', P' \ P' \ pre P\<^sub>2 \ \, \ \\<^bsub>/F\<^esub> P\<^sub>2 p2 Q',A' \ Q' \ Q \ A' \ A)" apply (induct rule: oghoare_induct, simp_all) apply blast apply (rename_tac \ \ F P c Q A) apply clarsimp apply(case_tac P, simp_all) apply clarsimp apply (rename_tac P'' Q'' A'' x) apply (rule_tac x="P''" in exI) apply (rule_tac x="Q''" in exI) apply clarsimp apply (rule conjI) apply (rule Conseq) apply (rule_tac x=P' in exI) apply fastforce apply (rule_tac x="A''" in exI) apply clarsimp apply fastforce done lemma oghoare_Cond[rule_format, OF _ refl]: "\, \ \\<^bsub>/F\<^esub> P x Q,A \ \c\<^sub>1 c\<^sub>2 b. x = (Cond b c\<^sub>1 c\<^sub>2) \ (\P' P\<^sub>1 P\<^sub>2 Q' A'. P = (AnnBin P' P\<^sub>1 P\<^sub>2) \ P' \ {s. (s\b \ s\pre P\<^sub>1) \ (s\b \ s\pre P\<^sub>2)} \ \, \ \\<^bsub>/F\<^esub> P\<^sub>1 c\<^sub>1 Q',A' \ \, \ \\<^bsub>/F\<^esub> P\<^sub>2 c\<^sub>2 Q',A' \ Q' \ Q \ A' \ A)" apply (induct rule: oghoare_induct, simp_all) apply (rule conjI) apply fastforce apply (rename_tac Q A P\<^sub>2 c\<^sub>2 r b) apply (rule_tac x=Q in exI) apply (rule_tac x=A in exI) apply fastforce apply (rename_tac \ \ F P c Q A) apply clarsimp apply (case_tac P, simp_all) apply fastforce done lemma oghoare_While[rule_format, OF _ refl]: "\, \ \\<^bsub>/F\<^esub> P x Q,A \ \ b c. x = While b c \ (\ r i P' A' Q'. P = AnnWhile r i P' \ \, \\\<^bsub>/F\<^esub> P' c i,A' \ r \ i \ i \ b \ pre P' \ i \ -b \ Q' \ Q' \ Q \ A' \ A)" apply (induct rule: oghoare_induct, simp_all) apply blast apply (rename_tac \ \ F P c Q A) apply clarsimp apply (rename_tac P' Q' A' b ca r i P'' A'' Q'') apply (case_tac P; simp) apply (rule_tac x= A'' in exI) apply (rule conjI) apply blast apply clarsimp apply (rule_tac x= "Q'" in exI) by fast lemma oghoare_Call: "\,\\\<^bsub>/F\<^esub> P x Q,A \ \p. x = Call p \ (\r n. P = (AnnCall r n) \ (\as P' f b. \ p = Some as \ (as ! n) = P' \ r \ pre P' \ \ p = Some b \ n < length as \ \,\ \\<^bsub>/F\<^esub> P' b Q,A))" apply (induct rule: oghoare_induct, simp_all) apply (rename_tac \ \ F P c Q A) apply clarsimp apply (case_tac P, simp_all) apply clarsimp apply (rule Conseq) apply (rule_tac x="{}" in exI) apply force done lemma oghoare_Parallel[rule_format, OF _ refl]: "\, \\\<^bsub>/F\<^esub> P x Q,A \ \cs. x = Parallel cs \ (\as. P = AnnPar as \ length as = length cs \ \(set (map postcond as)) \ Q \ \(set (map abrcond as)) \ A \ (\iQ' A'. \,\\\<^bsub>/F\<^esub> (pres (as!i)) (cs!i) Q', A' \ Q' \ postcond (as!i) \ A' \ abrcond (as!i)) \ interfree \ \ F as cs)" apply (induct rule: oghoare_induct, simp_all) apply clarsimp apply (drule_tac x=i in spec) apply fastforce apply clarsimp apply (case_tac P, simp_all) apply blast done lemma ann_matches_weaken[OF _ refl]: " ann_matches \ \ X c \ X = (weaken_pre P P') \ ann_matches \ \ P c" apply (induct arbitrary: P P' rule: ann_matches.induct) apply (case_tac P, simp_all, fastforce simp add: ann_matches.intros)+ done lemma oghoare_seq_imp_ann_matches: " \,\\\<^bsub>/F\<^esub> P a c Q,A \ ann_matches \ \ a c" apply (induct rule: oghoare_seq_induct, simp_all add: ann_matches.intros) apply (clarsimp, erule ann_matches_weaken)+ done lemma oghoare_imp_ann_matches: " \,\\\<^bsub>/F\<^esub> a c Q,A \ ann_matches \ \ a c" apply (induct rule: oghoare_induct, simp_all add: ann_matches.intros oghoare_seq_imp_ann_matches) apply (clarsimp, erule ann_matches_weaken)+ done (* intros *) lemma SkipRule: "P \ Q \ \, \ \\<^bsub>/F\<^esub> (AnnExpr P) Skip Q, A" apply (rule Conseq, simp) apply (rule exI, rule exI, rule exI) apply (rule conjI, rule Skip, auto) done lemma ThrowRule: "P \ A \ \, \ \\<^bsub>/F\<^esub> (AnnExpr P) Throw Q, A" apply (rule Conseq, simp) apply (rule exI, rule exI, rule exI) apply (rule conjI, rule Throw, auto) done lemma BasicRule: "P \ {s. (f s) \ Q} \ \, \ \\<^bsub>/F\<^esub> (AnnExpr P) (Basic f) Q, A" apply (rule Conseq, simp, rule exI[where x="{s. (f s) \ Q}"]) apply (rule exI[where x=Q], rule exI[where x=A]) apply simp apply (subgoal_tac "(P \ {s. f s \ Q}) = {s. f s \ Q}") apply (auto intro: Basic) done lemma SpecRule: "P \ {s. (\t. (s, t) \ r \ t \ Q) \ (\t. (s, t) \ r)} \ \, \ \\<^bsub>/F\<^esub> (AnnExpr P) (Spec r) Q, A" apply (rule Conseq, simp, rule exI[where x="{s. (\t. (s, t) \ r \ t \ Q) \ (\t. (s, t) \ r) }"]) apply (rule exI[where x=Q], rule exI[where x=A]) apply simp apply (subgoal_tac "(P \ {s. (\t. (s, t) \ r \ t \ Q) \ (\t. (s, t) \ r)}) = {s. (\t. (s, t) \ r \ t \ Q) \ (\t. (s, t) \ r)}") apply (auto intro: Spec) done lemma CondRule: "\ P \ {s. (s\b \ s\pre P\<^sub>1) \ (s\b \ s\pre P\<^sub>2)}; \, \ \\<^bsub>/F\<^esub> P\<^sub>1 c\<^sub>1 Q,A; \, \ \\<^bsub>/F\<^esub> P\<^sub>2 c\<^sub>2 Q,A \ \ \, \ \\<^bsub>/F\<^esub> (AnnBin P P\<^sub>1 P\<^sub>2) (Cond b c\<^sub>1 c\<^sub>2) Q,A" by (auto intro: Cond) lemma WhileRule: "\ r \ I; I \ b \ pre P; (I \ -b) \ Q; \, \ \\<^bsub>/F\<^esub> P c I,A \ \ \, \ \\<^bsub>/F\<^esub> (AnnWhile r I P) (While b c) Q,A" by (simp add: While) lemma AwaitRule: "\atom_com c ; \, \ \\<^bsub>/F\<^esub>P a c Q,A ; (r \ b) \ P\ \ \, \ \\<^bsub>/F\<^esub> (AnnRec r a) (Await b c) Q,A" apply (erule Await[rotated]) apply (erule (1) SeqConseq, simp+) done lemma rtranclp_1n_induct [consumes 1, case_names base step]: "\r\<^sup>*\<^sup>* a b; P a; \y z. \r y z; r\<^sup>*\<^sup>* z b; P y\ \ P z\ \ P b" by (induct rule: rtranclp.induct) (simp add: rtranclp.rtrancl_into_rtrancl)+ lemmas rtranclp_1n_induct2[consumes 1, case_names base step] = rtranclp_1n_induct[of _ "(ax,ay)" "(bx,by)", split_rule] lemma oghoare_seq_valid: " \\\<^bsub>/F\<^esub> P c\<^sub>1 R,A \ \\\<^bsub>/F\<^esub> R c\<^sub>2 Q,A \ \\\<^bsub>/F\<^esub> P Seq c\<^sub>1 c\<^sub>2 Q,A" apply (clarsimp simp add: valid_def) apply (rename_tac t c' s) apply (case_tac t) apply simp apply (drule (1) Seq_decomp_star) apply (erule disjE) apply fastforce apply clarsimp apply (rename_tac s1 t') apply (drule_tac x="Normal s" and y="Normal t'" in spec2) apply (erule_tac x="Skip" in allE) apply (fastforce simp: final_def) apply (clarsimp simp add: final_def) apply (drule Seq_decomp_star_Fault) apply (erule disjE) apply (rename_tac s2) apply (drule_tac x="Normal s" and y="Fault s2" in spec2) apply (erule_tac x="Skip" in allE) apply fastforce apply clarsimp apply (rename_tac s s2 s') apply (drule_tac x="Normal s" and y="Normal s'" in spec2) apply (erule_tac x="Skip" in allE) apply clarsimp apply (drule_tac x="Normal s'" and y="Fault s2" in spec2) apply (erule_tac x="Skip" in allE) apply clarsimp apply clarsimp apply (simp add: final_def) apply (drule Seq_decomp_star_Stuck) apply (erule disjE) apply fastforce apply clarsimp apply fastforce done lemma oghoare_if_valid: "\\\<^bsub>/F\<^esub> P\<^sub>1 c\<^sub>1 Q,A \ \\\<^bsub>/F\<^esub> P\<^sub>2 c\<^sub>2 Q,A \ r \ b \ P\<^sub>1 \ r \ - b \ P\<^sub>2 \ \\\<^bsub>/F\<^esub> r Cond b c\<^sub>1 c\<^sub>2 Q,A" apply (simp (no_asm) add: valid_def) apply (clarsimp) apply (erule converse_rtranclpE) apply (clarsimp simp: final_def) apply (erule step_Normal_elim_cases) apply (fastforce simp: valid_def[where c=c\<^sub>1]) apply (fastforce simp: valid_def[where c=c\<^sub>2]) done lemma Skip_normal_steps_end: "\ \ (Skip, Normal s) \\<^sup>* (c, s') \ c = Skip \ s' = Normal s" apply (erule converse_rtranclpE) apply simp apply (erule step_Normal_elim_cases) done lemma Throw_normal_steps_end: "\ \ (Throw, Normal s) \\<^sup>* (c, s') \ c = Throw \ s' = Normal s" apply (erule converse_rtranclpE) apply simp apply (erule step_Normal_elim_cases) done lemma while_relpower_induct: "\t c' x . \\\<^bsub>/F\<^esub> P c i,A \ i \ b \ P \ i \ - b \ Q \ final (c', t) \ x \ i \ t \ Fault ` F \ c' = Throw \ t \ Normal ` A \ (step \ ^^ n) (While b c, Normal x) (c', t) \ c' = Skip \ t \ Normal ` Q" apply (induct n rule:less_induct) apply (rename_tac n t c' x) apply (case_tac n) apply (clarsimp simp: final_def) apply clarify apply (simp only: relpowp.simps) apply (subst (asm) relpowp_commute[symmetric]) apply clarsimp apply (erule step_Normal_elim_cases) apply clarsimp apply (rename_tac t c' x v) apply(case_tac "t") apply clarsimp apply(drule Seq_decomp_relpow) apply(simp add: final_def) apply(erule disjE, erule conjE) apply clarify apply(drule relpowp_imp_rtranclp) apply (simp add: valid_def) apply (rename_tac x n t' n1) apply (drule_tac x="Normal x" in spec) apply (drule_tac x="Normal t'" in spec) apply (drule spec[where x=Throw]) apply (fastforce simp add: final_def) apply clarsimp apply (rename_tac c' x n t' t n1 n2) apply (drule_tac x=n2 and y="Normal t'" in meta_spec2) apply (drule_tac x=c' and y="t" in meta_spec2) apply (erule meta_impE, fastforce) apply (erule meta_impE, fastforce) apply (erule meta_impE) apply(drule relpowp_imp_rtranclp) apply (simp add: valid_def) apply (drule_tac x="Normal x" and y="Normal t" in spec2) apply (drule spec[where x=Skip]) apply (fastforce simp add: final_def) apply assumption apply clarsimp apply (rename_tac c' s n f) apply (subgoal_tac "c' = Skip", simp) prefer 2 apply (case_tac c'; fastforce simp: final_def) apply (drule Seq_decomp_relpowp_Fault) apply (erule disjE) apply (clarsimp simp: valid_def) apply (drule_tac x="Normal s" and y="Fault f" in spec2) apply (drule spec[where x=Skip]) apply(drule relpowp_imp_rtranclp) apply (fastforce simp: final_def) apply clarsimp apply (rename_tac t n1 n2) apply (subgoal_tac "t \ i") prefer 2 apply (fastforce dest:relpowp_imp_rtranclp simp: final_def valid_def) apply (drule_tac x=n2 in meta_spec) apply (drule_tac x="Fault f" in meta_spec) apply (drule meta_spec[where x=Skip]) apply (drule_tac x=t in meta_spec) apply (fastforce simp: final_def) apply clarsimp apply (rename_tac c' s n) apply (subgoal_tac "c' = Skip", simp) prefer 2 apply (case_tac c'; fastforce simp: final_def) apply (drule Seq_decomp_relpowp_Stuck) apply clarsimp apply (erule disjE) apply (simp add:valid_def) apply (drule_tac x="Normal s" and y="Stuck" in spec2) apply clarsimp apply (drule spec[where x=Skip]) apply(drule relpowp_imp_rtranclp) apply (fastforce) apply clarsimp apply (rename_tac t n1 n2) apply (subgoal_tac "t \ i") prefer 2 apply (fastforce dest:relpowp_imp_rtranclp simp: final_def valid_def) apply (drule_tac x=n2 in meta_spec) apply (drule meta_spec[where x=Stuck]) apply (drule meta_spec[where x=Skip]) apply (drule_tac x=t in meta_spec) apply (fastforce simp: final_def) apply clarsimp apply (drule relpowp_imp_rtranclp) apply (drule Skip_normal_steps_end) apply fastforce done lemma valid_weaken_pre: "\\\<^bsub>/F\<^esub> P c Q,A \ P' \ P \ \\\<^bsub>/F\<^esub> P' c Q,A" by (fastforce simp: valid_def) lemma valid_strengthen_post: "\\\<^bsub>/F\<^esub> P c Q,A \ Q \ Q' \ \\\<^bsub>/F\<^esub> P c Q',A" by (fastforce simp: valid_def) lemma valid_strengthen_abr: "\\\<^bsub>/F\<^esub> P c Q,A \ A \ A' \ \\\<^bsub>/F\<^esub> P c Q,A'" by (fastforce simp: valid_def) lemma oghoare_while_valid: "\\\<^bsub>/F\<^esub> P c i,A \ i \ b \ P \ i \ - b \ Q \ \\\<^bsub>/F\<^esub> i While b c Q,A" apply (simp (no_asm) add: valid_def) apply (clarsimp simp add: ) apply (drule rtranclp_imp_relpowp) apply (clarsimp) by (erule while_relpower_induct) lemma oghoare_catch_valid: "\\\<^bsub>/F\<^esub> P\<^sub>1 c\<^sub>1 Q,P\<^sub>2 \ \\\<^bsub>/F\<^esub> P\<^sub>2 c\<^sub>2 Q,A \ \\\<^bsub>/F\<^esub> P\<^sub>1 Catch c\<^sub>1 c\<^sub>2 Q,A" apply (clarsimp simp add: valid_def[where c="Catch _ _"]) apply (rename_tac t c' s) apply (case_tac t) apply simp apply (drule Catch_decomp_star) apply (fastforce simp: final_def) apply clarsimp apply (erule disjE) apply (clarsimp simp add: valid_def[where c="c\<^sub>1"]) apply (rename_tac s x t) apply (drule_tac x="Normal s" in spec) apply (drule_tac x="Normal t" in spec) apply (drule_tac x=Throw in spec) apply (fastforce simp: final_def valid_def) apply (clarsimp simp add: valid_def[where c="c\<^sub>1"]) apply (rename_tac s t) apply (drule_tac x="Normal s" in spec) apply (drule_tac x="Normal t" in spec) apply (drule_tac x=Skip in spec) apply (fastforce simp: final_def) apply (rename_tac c' s t) apply (simp add: final_def) apply (drule Catch_decomp_star_Fault) apply clarsimp apply (erule disjE) apply (clarsimp simp: valid_def[where c=c\<^sub>1] final_def) apply (fastforce simp: valid_def final_def) apply (simp add: final_def) apply (drule Catch_decomp_star_Stuck) apply clarsimp apply (erule disjE) apply (clarsimp simp: valid_def[where c=c\<^sub>1] final_def) apply (fastforce simp: valid_def final_def) done lemma ann_matches_imp_assertionsR: "ann_matches \ \ a c \ \ pre_par a \ assertionsR \ \ Q A a c (pre a)" by (induct arbitrary: Q A rule: ann_matches.induct , (fastforce intro: assertionsR.intros )+) lemma ann_matches_imp_assertionsR': "ann_matches \ \ a c \ a' \ pre_set a \ assertionsR \ \ Q A a c a'" apply (induct arbitrary: Q A rule: ann_matches.induct) apply ((fastforce intro: assertionsR.intros )+)[12] apply simp apply (erule bexE) apply (simp only: in_set_conv_nth) apply (erule exE) apply (drule_tac x=i in spec) apply clarsimp apply (erule AsParallelExprs) apply simp done lemma rtranclp_conjD: "(\x1 x2. r1 x1 x2 \ r2 x1 x2)\<^sup>*\<^sup>* x1 x2 \ r1\<^sup>*\<^sup>* x1 x2 \ r2\<^sup>*\<^sup>* x1 x2" by (metis (no_types, lifting) rtrancl_mono_proof) lemma rtranclp_mono' : "r\<^sup>*\<^sup>* a b \ r \ s \ s\<^sup>*\<^sup>* a b" by (metis rtrancl_mono_proof sup.orderE sup2CI) lemma state_upd_in_atomicsR[rule_format, OF _ refl refl]: "\\ cf \ cf' \ cf = (c, Normal s) \ cf' = (c', Normal t) \ s \ t \ ann_matches \ \ a c \ s \ pre a \ (\p cm x. atomicsR \ \ a c (p, cm) \ s \ p \ \ \ (cm, Normal s) \ (x, Normal t) \ final (x, Normal t))" apply (induct arbitrary: c c' s t a rule: step.induct, simp_all) apply clarsimp apply (erule ann_matches.cases, simp_all) apply (rule exI)+ apply (rule conjI) apply (rule atomicsR.intros) apply clarsimp apply (rule_tac x="Skip" in exI) apply (simp add: final_def) apply (rule step.Basic) apply clarsimp apply (erule ann_matches.cases, simp_all) apply (rule exI)+ apply (rule conjI) apply (rule atomicsR.intros) apply clarsimp apply (rule_tac x="Skip" in exI) apply (simp add: final_def) apply (erule step.Spec) apply clarsimp apply (erule ann_matches.cases, simp_all) apply clarsimp apply (drule meta_spec)+ apply (erule meta_impE, rule conjI, (rule refl)+)+ apply clarsimp apply (erule (1) meta_impE) apply (erule meta_impE, fastforce) apply clarsimp apply (rule exI)+ apply (rule conjI) apply (erule AtSeqExpr1) apply fastforce apply clarsimp apply (erule ann_matches.cases, simp_all) apply clarsimp apply (drule meta_spec)+ apply (erule meta_impE, rule conjI, (rule refl)+)+ apply clarsimp apply (erule (1) meta_impE) apply (erule meta_impE, fastforce) apply clarsimp apply (rule exI)+ apply (rule conjI) apply (erule AtCatchExpr1) apply fastforce apply (erule ann_matches.cases, simp_all) apply clarsimp apply (drule meta_spec)+ apply (erule meta_impE, rule conjI, (rule refl)+)+ apply clarsimp apply (erule meta_impE) apply fastforce apply (erule meta_impE) apply (case_tac "i=0"; fastforce) apply clarsimp apply (rule exI)+ apply (rule conjI) apply (erule AtParallelExprs) apply fastforce apply (drule_tac x=i in spec) apply clarsimp apply fastforce apply (erule ann_matches.cases, simp_all) apply clarsimp apply (rule exI)+ apply (rule conjI) apply (rule AtAwait) apply clarsimp apply (rename_tac c' sa t aa e r ba) apply (rule_tac x=c' in exI) apply (rule conjI) apply (erule step.Await) apply (erule rtranclp_mono') apply clarsimp apply assumption+ apply (simp add: final_def) done lemma oghoare_atom_com_sound: "\, \ \\<^bsub>/F\<^esub>P a c Q,A \ atom_com c \ \ \\<^bsub>/F\<^esub> P c Q, A" unfolding valid_def proof (induct rule: oghoare_seq_induct) case SeqSkip thus ?case by (fastforce elim: converse_rtranclpE step_Normal_elim_cases(1)) next case SeqThrow thus ?case by (fastforce elim: converse_rtranclpE step_Normal_elim_cases) next case SeqBasic thus ?case by (fastforce elim: converse_rtranclpE step_Normal_elim_cases simp: final_def) next case (SeqSpec \ \ F r Q A) thus ?case apply clarsimp apply (erule converse_rtranclpE) apply (clarsimp simp: final_def) apply (erule step_Normal_elim_cases) apply (fastforce elim!: converse_rtranclpE step_Normal_elim_cases) by clarsimp next case (SeqSeq \ \ F P\<^sub>1 c\<^sub>1 P\<^sub>2 A c\<^sub>2 Q) show ?case using SeqSeq by (fold valid_def) (fastforce intro: oghoare_seq_valid simp: valid_weaken_pre) next case (SeqCatch \ \ F P\<^sub>1 c\<^sub>1 Q P\<^sub>2 c\<^sub>2 A) thus ?case apply (fold valid_def) apply simp apply (fastforce elim: oghoare_catch_valid)+ done next case (SeqCond \ \ F P b c1 Q A c2) thus ?case by (fold valid_def) (fastforce intro:oghoare_if_valid) next case (SeqWhile \ \ F P c A b) thus ?case by (fold valid_def) (fastforce elim: valid_weaken_pre[rotated] oghoare_while_valid) next case (SeqGuard \ \ F P c Q A r g f) thus ?case apply (fold valid_def) apply (simp (no_asm) add: valid_def) apply clarsimp apply (erule converse_rtranclpE) apply (fastforce simp: final_def) apply clarsimp apply (erule step_Normal_elim_cases) apply (case_tac "r \ - g \ {}") apply clarsimp apply (fastforce simp: valid_def) apply clarsimp apply (fastforce simp: valid_def) apply clarsimp apply (case_tac "r \ - g \ {}") apply (fastforce dest: no_steps_final simp:final_def) apply (fastforce dest: no_steps_final simp:final_def) done next case (SeqCall \ p f \ F P Q A) thus ?case by simp next case (SeqDynCom r fa \ \ F P c Q A) thus ?case apply - apply clarsimp apply (erule converse_rtranclpE) apply (clarsimp simp: final_def) apply clarsimp apply (erule step_Normal_elim_cases) apply clarsimp apply (rename_tac t c' x) apply (drule_tac x=x in spec) apply (drule_tac x=x in bspec, fastforce) apply clarsimp apply (drule_tac x="Normal x" in spec) apply (drule_tac x="t" in spec) apply (drule_tac x="c'" in spec) apply fastforce+ done next case (SeqConseq \ \ F P c Q A) thus ?case apply (clarsimp) apply (rename_tac t c' x) apply (erule_tac x="Normal x" in allE) apply (erule_tac x="t" in allE) apply (erule_tac x="c'" in allE) apply (clarsimp simp: pre_weaken_pre) apply force done qed simp_all lemma ParallelRuleAnn: " length as = length cs \ \i,\ \\<^bsub>/F \<^esub>(pres (as ! i)) (cs ! i) (postcond (as ! i)),(abrcond (as ! i)) \ interfree \ \ F as cs \ \(set (map postcond as)) \ Q \ \(set (map abrcond as)) \ A \ \,\ \\<^bsub>/F \<^esub>(AnnPar as) (Parallel cs) Q,A" apply (erule (3) Parallel) apply auto done lemma oghoare_step[rule_format, OF _ refl refl]: shows "\ \ cf \ cf' \ cf = (c, Normal s) \ cf' = (c', Normal t) \ \,\\\<^bsub>/F \<^esub>a c Q,A \ s \ pre a \ \a'. \,\\\<^bsub>/F \<^esub>a' c' Q,A \ t \ pre a' \ (\as. assertionsR \ \ Q A a' c' as \ assertionsR \ \ Q A a c as) \ (\pm cm. atomicsR \ \ a' c' (pm, cm) \ atomicsR \ \ a c (pm, cm))" proof (induct arbitrary:c c' s a t Q A rule: step.induct) case (Parallel i cs s c' s' ca c'a sa a t Q A) thus ?case apply (clarsimp simp:) apply (drule oghoare_Parallel) apply clarsimp apply (rename_tac as) apply (frule_tac x=i in spec, erule (1) impE) apply (elim exE conjE) apply (drule meta_spec)+ apply (erule meta_impE, rule conjI, (rule refl)+)+ apply (erule meta_impE) apply (rule_tac P="(pres (as ! i))" in Conseq) apply (rule exI[where x="{}"]) apply (rule_tac x="Q'" in exI) apply (rule_tac x="A'" in exI) apply (simp) apply (erule meta_impE, simp) apply clarsimp apply (rule_tac x="AnnPar (as[i:=(a',postcond(as!i), abrcond(as!i))])" in exI) apply (rule conjI) apply (rule ParallelRuleAnn, simp) apply clarsimp apply (rename_tac j) apply (drule_tac x=j in spec) apply clarsimp apply (case_tac "i = j") apply (clarsimp simp: ) apply (rule Conseq) apply (rule exI[where x="{}"]) apply (fastforce) apply (simp add: ) apply (clarsimp simp: interfree_def) apply (rename_tac i' j') apply (drule_tac x=i' and y=j' in spec2) apply clarsimp apply (case_tac "i = i'") apply clarsimp apply (simp add: interfree_aux_def prod.case_eq_if ) apply clarsimp apply (case_tac "j' = i") apply (clarsimp) apply (simp add: interfree_aux_def prod.case_eq_if) apply clarsimp apply (clarsimp) apply (erule subsetD) apply (clarsimp simp: in_set_conv_nth) apply (rename_tac a' x a b c i') apply (case_tac "i' = i") apply clarsimp apply (drule_tac x="(a', b, c)" in bspec, simp) apply (fastforce simp add: in_set_conv_nth) apply fastforce apply (drule_tac x="(a, b, c)" in bspec, simp) apply (simp add: in_set_conv_nth) apply (rule_tac x=i' in exI) apply clarsimp apply fastforce apply clarsimp apply (erule_tac A="(\x\set as. abrcond x) " in subsetD) apply (clarsimp simp: in_set_conv_nth) apply (rename_tac a b c j) apply (case_tac "j = i") apply clarsimp apply (rule_tac x="as!i" in bexI) apply simp apply clarsimp apply clarsimp apply (rule_tac x="(a,b,c)" in bexI) apply simp apply (clarsimp simp:in_set_conv_nth) apply (rule_tac x=j in exI) apply fastforce apply (rule conjI) apply (case_tac "s = Normal t") apply clarsimp apply (clarsimp simp: in_set_conv_nth) apply (rename_tac a b c j) apply (case_tac "j = i") apply clarsimp apply clarsimp apply (drule_tac x="as!j" in bspec) apply (clarsimp simp add: in_set_conv_nth) apply (rule_tac x=j in exI) apply fastforce apply clarsimp apply (frule state_upd_in_atomicsR, simp) apply (erule oghoare_imp_ann_matches) apply (clarsimp simp: in_set_conv_nth) apply fastforce apply (clarsimp simp: in_set_conv_nth) apply (rename_tac j) apply (case_tac "j = i") apply clarsimp apply clarsimp apply (thin_tac "\,\ \\<^bsub>/F \<^esub>a' c' (postcond (as ! i)),(abrcond (as ! i))") apply (simp add: interfree_def interfree_aux_def) apply (drule_tac x="j" and y=i in spec2) apply (simp add: prod.case_eq_if) apply (drule spec2, drule (1) mp) apply clarsimp apply (case_tac "pre_par a") apply (subst pre_set) apply clarsimp apply (drule_tac x="as!j" in bspec) apply (clarsimp simp: in_set_conv_nth) apply (rule_tac x=j in exI) apply fastforce apply clarsimp apply (frule (1) pre_imp_pre_set) apply (rename_tac as Q' A' a' a b c p cm x j X) apply (drule_tac x="X" in spec, erule_tac P="assertionsR \ \ b c a (cs ! j) X" in impE) apply (rule ann_matches_imp_assertionsR') apply (drule_tac x=j in spec, clarsimp) apply (erule (1) oghoare_imp_ann_matches) apply (rename_tac a b c p cm x j X ) apply (thin_tac "\\\<^bsub>/F\<^esub> (b \ p) cm b,b") apply (thin_tac " \\\<^bsub>/F\<^esub> (c \ p) cm c,c") apply (simp add: valid_def) apply (drule_tac x="Normal sa" in spec) apply (drule_tac x="Normal t" in spec) apply (drule_tac x=x in spec) apply (erule impE, fastforce) apply force apply (drule_tac x=j in spec) apply clarsimp apply (rename_tac a b c p cm x j Q'' A'') apply (drule_tac x="pre a" in spec,erule impE, rule ann_matches_imp_assertionsR) apply (erule (1) oghoare_imp_ann_matches) apply (thin_tac " \\\<^bsub>/F\<^esub> (b \ p) cm b,b") apply (thin_tac " \\\<^bsub>/F\<^esub> (c \ p) cm c,c") apply (simp add: valid_def) apply (drule_tac x="Normal sa" in spec) apply (drule_tac x="Normal t" in spec) apply (drule_tac x=x in spec) apply (erule impE, fastforce) apply clarsimp apply (drule_tac x="as ! j" in bspec) apply (clarsimp simp: in_set_conv_nth) apply (rule_tac x=j in exI, fastforce) apply clarsimp apply fastforce apply (rule conjI) apply (clarsimp simp: ) apply (erule assertionsR.cases ; simp) apply (clarsimp simp: ) apply (rename_tac j a) apply (case_tac "j = i") apply clarsimp apply (drule_tac x=a in spec, erule (1) impE) apply (erule (1) AsParallelExprs) apply (subst (asm) nth_list_update_neq, simp) apply (erule_tac i=j in AsParallelExprs) apply fastforce apply clarsimp apply (rule AsParallelSkips) apply (clarsimp simp:) apply (rule equalityI) apply (clarsimp simp: in_set_conv_nth) apply (rename_tac a' x a b c j) apply (case_tac "j = i") apply (thin_tac "\a\set as. sa \ precond a") apply clarsimp apply (drule_tac x="(a', b, c)" in bspec) apply (clarsimp simp: in_set_conv_nth) apply (rule_tac x="i" in exI) apply fastforce apply fastforce apply (drule_tac x="as ! j" in bspec) apply (clarsimp simp: in_set_conv_nth) apply (rule_tac x=j in exI) apply fastforce apply clarsimp apply (drule_tac x=" as ! j" in bspec) apply (clarsimp simp: in_set_conv_nth) apply (rule_tac x=j in exI, fastforce) apply fastforce apply (clarsimp simp: in_set_conv_nth) apply (rename_tac x a b c j) apply (thin_tac "\a\set as. sa \ precond a") apply (case_tac "j = i") apply clarsimp apply (drule_tac x="as!i" in bspec, fastforce) apply fastforce apply clarsimp apply (drule_tac x="as!j" in bspec) apply (clarsimp simp: in_set_conv_nth) apply (rule_tac x=j in exI, fastforce) apply fastforce apply clarsimp apply (erule atomicsR.cases ; simp) apply clarsimp apply (rename_tac j atc atp) apply (case_tac "j = i") apply clarsimp apply (drule_tac x=atc and y=atp in spec2, erule impE) apply (clarsimp) apply (erule (1) AtParallelExprs) apply (subst (asm) nth_list_update_neq, simp) apply (erule_tac i=j in AtParallelExprs) apply (clarsimp) done next case (Basic f s c c' sa a t Q A) thus ?case apply clarsimp apply (drule oghoare_Basic) apply clarsimp apply (rule_tac x="AnnExpr Q" in exI) apply clarsimp apply (rule conjI) apply (rule SkipRule) apply fastforce apply (rule conjI) apply fastforce apply clarsimp apply (drule assertionsR.cases, simp_all) apply (rule assertionsR.AsBasicSkip) done next case (Spec s t r c c' sa a ta Q A) thus ?case apply clarsimp apply (drule oghoare_Spec) apply clarsimp apply (rule_tac x="AnnExpr Q" in exI) apply clarsimp apply (rule conjI) apply (rule SkipRule) apply fastforce apply (rule conjI) apply force apply clarsimp apply (erule assertionsR.cases, simp_all) apply clarsimp apply (rule assertionsR.AsSpecSkip) done next case (Guard s g f c ca c' sa a t Q A) thus ?case apply - apply clarsimp apply (drule oghoare_Guard) apply clarsimp apply (rule exI, rule conjI, assumption) by (fastforce dest: oghoare_Guard intro:assertionsR.intros atomicsR.intros) next case (GuardFault s g f c ca c' sa a t Q A) thus ?case by (fastforce dest: oghoare_Guard intro:assertionsR.intros atomicsR.intros) next case (Seq c\<^sub>1 s c\<^sub>1' s' c\<^sub>2 c c' sa a t A Q) thus ?case apply (clarsimp simp:) apply (drule oghoare_Seq) apply clarsimp apply (drule meta_spec)+ apply (erule meta_impE, rule conjI, (rule refl)+)+ apply (erule meta_impE) apply (rule Conseq) apply (rule exI[where x="{}"]) apply (rule exI)+ apply (rule conjI) apply (simp) apply (erule (1) conjI) apply clarsimp apply (rule_tac x="AnnComp a' P\<^sub>2" in exI, rule conjI) apply (rule oghoare_oghoare_seq.Seq) apply (rule Conseq) apply (rule_tac x="{}" in exI) apply (fastforce) apply (rule Conseq) apply (rule_tac x="{}" in exI) apply (fastforce) apply clarsimp apply (rule conjI) apply clarsimp apply (erule assertionsR.cases, simp_all) apply clarsimp apply (drule_tac x=a in spec, simp) apply (erule AsSeqExpr1) apply clarsimp apply (erule AsSeqExpr2) apply clarsimp apply (erule atomicsR.cases, simp_all) apply clarsimp apply (drule_tac x="a" and y=b in spec2, simp) apply (erule AtSeqExpr1) apply clarsimp apply (erule AtSeqExpr2) done next case (SeqSkip c\<^sub>2 s c c' sa a t Q A) thus ?case apply clarsimp apply (drule oghoare_Seq) apply clarsimp apply (rename_tac P\<^sub>1 P\<^sub>2 P' Q' A') apply (rule_tac x=P\<^sub>2 in exI) apply (rule conjI, rule Conseq) apply (rule_tac x="{}" in exI) apply (fastforce) apply (rule conjI) apply (drule oghoare_Skip) apply fastforce apply (rule conjI) apply clarsimp apply (erule assertionsR.AsSeqExpr2) apply clarsimp apply (fastforce intro: atomicsR.intros) done next case (SeqThrow c\<^sub>2 s c c' sa a t Q A) thus ?case apply clarsimp apply (drule oghoare_Seq) apply clarsimp apply (rename_tac P\<^sub>1 P\<^sub>2 P' Q' A') apply (rule_tac x=P\<^sub>1 in exI) apply (drule oghoare_Throw) apply clarsimp apply (rename_tac P'') apply (rule conjI, rule Conseq) apply (rule_tac x="{}" in exI) apply (rule_tac x="Q'" in exI) apply (rule_tac x="P''" in exI) apply (fastforce intro: Throw) apply clarsimp apply (erule assertionsR.cases, simp_all) apply clarsimp apply (rule AsSeqExpr1) apply (rule AsThrow) done next case (CondTrue s b c\<^sub>1 c\<^sub>2 c sa c' s' ann) thus ?case apply (clarsimp) apply (drule oghoare_Cond) apply clarsimp apply (rename_tac P' P\<^sub>1 P\<^sub>2 Q' A') apply (rule_tac x= P\<^sub>1 in exI) apply (rule conjI) apply (rule Conseq, rule_tac x="{}" in exI, fastforce) apply (rule conjI, fastforce) apply (rule conjI) apply (fastforce elim: assertionsR.cases intro: AsCondExpr1) apply (fastforce elim: atomicsR.cases intro: AtCondExpr1) done next case (CondFalse s b c\<^sub>1 c\<^sub>2 c sa c' s' ann) thus ?case apply (clarsimp) apply (drule oghoare_Cond) apply clarsimp apply (rename_tac P' P\<^sub>1 P\<^sub>2 Q' A') apply (rule_tac x= P\<^sub>2 in exI) apply (rule conjI) apply (rule Conseq, rule_tac x="{}" in exI, fastforce) apply (rule conjI, fastforce) apply (rule conjI) apply (fastforce elim: assertionsR.cases intro: AsCondExpr2) apply (fastforce elim: atomicsR.cases intro: AtCondExpr2) done next case (WhileTrue s b c ca sa c' s' ann) thus ?case apply clarsimp apply (frule oghoare_While) apply clarsimp apply (rename_tac r i P' A' Q') apply (rule_tac x="AnnComp P' (AnnWhile i i P')" in exI) apply (rule conjI) apply (rule Seq) apply (rule Conseq) apply (rule_tac x="{}" in exI) apply (rule_tac x="i" in exI) apply (rule_tac x="A'" in exI) apply (subst weaken_pre_empty) apply clarsimp apply (rule While) apply (rule Conseq) apply (rule_tac x="{}" in exI) apply (rule_tac x="i" in exI) apply (rule_tac x="A'" in exI) apply (subst weaken_pre_empty) apply clarsimp apply clarsimp apply force apply simp apply simp apply (rule conjI) apply blast apply (rule conjI) apply clarsimp apply (erule assertionsR.cases, simp_all) apply clarsimp apply (rule AsWhileExpr) apply clarsimp apply (erule assertionsR.cases,simp_all) apply clarsimp apply (erule AsWhileExpr) apply clarsimp apply (rule AsWhileInv) apply clarsimp apply (rule AsWhileInv) apply clarsimp apply (rule AsWhileSkip) apply clarsimp apply (erule atomicsR.cases, simp_all) apply clarsimp apply (rule AtWhileExpr) apply clarsimp+ apply (erule atomicsR.cases, simp_all) apply clarsimp apply (erule AtWhileExpr) done next case (WhileFalse s b c ca sa c' ann s' Q A) thus ?case apply clarsimp apply (drule oghoare_While) apply clarsimp apply (rule_tac x="AnnExpr Q" in exI) apply (rule conjI) apply (rule SkipRule) apply blast apply (rule conjI) apply fastforce apply clarsimp apply (erule assertionsR.cases, simp_all) apply (drule sym, simp, rule AsWhileSkip) done next case (Call p bs s c c' sa a t Q A) thus ?case apply clarsimp apply (drule oghoare_Call) apply clarsimp apply (rename_tac n as) apply (rule_tac x="as ! n" in exI) apply clarsimp apply (rule conjI, fastforce) apply (rule conjI) apply clarsimp apply (erule (2) AsCallExpr) apply fastforce apply clarsimp apply (erule (2) AtCallExpr) apply simp done next case (DynCom c s ca c' sa a t Q A) thus ?case apply - apply clarsimp apply (drule oghoare_DynCom) apply clarsimp apply (drule_tac x=t in bspec, assumption) apply (rule exI) apply (erule conjI) apply (rule conjI, fastforce) apply (rule conjI) apply clarsimp apply (erule (1) AsDynComExpr) apply (clarsimp) apply (erule (1) AtDynCom) done next case (Catch c\<^sub>1 s c\<^sub>1' s' c\<^sub>2 c c' sa a t Q A) thus ?case apply (clarsimp simp:) apply (drule oghoare_Catch) apply clarsimp apply (drule meta_spec)+ apply (erule meta_impE, rule conjI, (rule refl)+)+ apply (erule meta_impE) apply (rule Conseq) apply (rule exI[where x="{}"]) apply (rule exI)+ apply (rule conjI) apply (simp) apply (erule (1) conjI) apply clarsimp apply (rename_tac P\<^sub>1 P\<^sub>2 P' Q' A' a') apply (rule_tac x="AnnComp a' P\<^sub>2" in exI, rule conjI) apply (rule oghoare_oghoare_seq.Catch) apply (rule Conseq) apply (rule_tac x="{}" in exI) apply (fastforce) apply (rule Conseq) apply (rule_tac x="{}" in exI) apply (fastforce) apply clarsimp apply (rule conjI) apply clarsimp apply (erule assertionsR.cases, simp_all) apply clarsimp apply (rename_tac a) apply (drule_tac x=a in spec, simp) apply (erule AsCatchExpr1) apply clarsimp apply (erule AsCatchExpr2) apply clarsimp apply (erule atomicsR.cases, simp_all) apply clarsimp apply (rename_tac a b a2) apply (drule_tac x="a" and y=b in spec2, simp) apply (erule AtCatchExpr1) apply clarsimp apply (erule AtCatchExpr2) done next case (CatchSkip c\<^sub>2 s c c' sa a t Q A) thus ?case apply clarsimp apply (drule oghoare_Catch, clarsimp) apply (rename_tac P\<^sub>1 P\<^sub>2 P' Q' A') apply (rule_tac x=P\<^sub>1 in exI) apply clarsimp apply (rule conjI) apply (rule Conseq) apply (rule_tac x="{}" in exI) apply (drule oghoare_Skip) apply clarsimp apply (rule_tac x=Q' in exI) apply (rule_tac x=A' in exI) apply (rule conjI, erule SkipRule) apply clarsimp apply clarsimp apply (rule AsCatchExpr1) apply (erule assertionsR.cases, simp_all) apply (rule assertionsR.AsSkip) done next case (CatchThrow c\<^sub>2 s c c' sa a t Q A) thus ?case apply clarsimp apply (drule oghoare_Catch, clarsimp) apply (rename_tac P\<^sub>1 P\<^sub>2 P' Q' A') apply (rule_tac x=P\<^sub>2 in exI) apply (rule conjI) apply (rule Conseq) apply (rule_tac x="{}" in exI) apply (fastforce ) apply (rule conjI) apply (drule oghoare_Throw) apply clarsimp apply fastforce apply (rule conjI) apply (clarsimp) apply (erule AsCatchExpr2) apply clarsimp apply (erule AtCatchExpr2) done next case (ParSkip cs s c c' sa a t Q A) thus ?case apply clarsimp apply (drule oghoare_Parallel) apply clarsimp apply (rename_tac as) apply (rule_tac x="AnnExpr (\x\set as. postcond x)" in exI) apply (rule conjI, rule SkipRule) apply blast apply (rule conjI) apply simp apply clarsimp apply (simp only: in_set_conv_nth) apply clarsimp apply (drule_tac x="i" in spec) apply clarsimp apply (drule_tac x="cs!i" in bspec) apply clarsimp apply clarsimp apply (drule oghoare_Skip) apply clarsimp apply (drule_tac x="as!i" in bspec) apply (clarsimp simp: in_set_conv_nth) apply (rule_tac x=i in exI, fastforce) apply clarsimp apply blast apply clarsimp apply (erule assertionsR.cases; simp) apply clarsimp apply (rule AsParallelSkips; clarsimp) done next case (ParThrow cs s c c' sa a t Q A) thus ?case apply clarsimp apply (drule oghoare_Parallel) apply (clarsimp simp: in_set_conv_nth) apply (drule_tac x=i in spec) apply clarsimp apply (drule oghoare_Throw) apply clarsimp apply (rename_tac i as Q' A' P') apply (rule_tac x="AnnExpr P'" in exI) apply (rule conjI) apply (rule ThrowRule) apply clarsimp apply (erule_tac A="(\x\set as. abrcond x)" in subsetD[where B=A], force) apply (rule conjI) apply (drule_tac x="as!i" in bspec) apply (clarsimp simp: in_set_conv_nth) apply (rule_tac x=i in exI, fastforce) apply (fastforce) apply clarsimp apply (erule AsParallelExprs) apply clarsimp apply (erule assertionsR.cases, simp_all) apply (rule AsThrow) done next case (Await x b c c' x' c'' c''' x'' a x''' Q A) thus ?case apply (clarsimp) apply (drule oghoare_Await) apply clarsimp apply (drule rtranclp_conjD) apply clarsimp apply (erule disjE) apply clarsimp apply (rename_tac P' Q' A') apply (rule_tac x="AnnExpr Q" in exI) apply clarsimp apply (rule conjI) apply (rule Skip) apply (rule conjI) apply (drule (1) oghoare_atom_com_sound) apply (fastforce simp: final_def valid_def) apply clarsimp apply (erule assertionsR.cases, simp_all) apply clarsimp apply (rule AsAwaitSkip) apply (rule_tac x="AnnExpr A" in exI) apply clarsimp apply (rule conjI) apply (rule Throw) apply (rule conjI) apply (drule (1) oghoare_atom_com_sound) apply (fastforce simp: final_def valid_def) apply clarsimp apply (erule assertionsR.cases, simp_all) apply clarsimp apply (rule AsAwaitThrow) done qed simp_all lemma oghoare_steps[rule_format, OF _ refl refl]: "\ \ cf \\<^sup>* cf' \ cf = (c, Normal s) \ cf' = (c', Normal t) \ \,\\\<^bsub>/F \<^esub>a c Q,A \ s \ pre a \ \a'. \,\\\<^bsub>/F \<^esub>a' c' Q,A \ t \ pre a' \ (\as. assertionsR \ \ Q A a' c' as \ assertionsR \ \ Q A a c as) \ (\pm cm. atomicsR \ \ a' c' (pm, cm) \ atomicsR \ \ a c (pm, cm))" apply (induct arbitrary: a c s c' t rule: converse_rtranclp_induct) apply fastforce apply clarsimp apply (frule Normal_pre_star) apply clarsimp apply (drule (2) oghoare_step) apply clarsimp apply ((drule meta_spec)+, (erule meta_impE, rule conjI, (rule refl)+)+) apply (erule (1) meta_impE)+ apply clarsimp apply (rule exI) apply (rule conjI, fastforce)+ apply force done lemma oghoare_sound_Parallel_Normal_case[rule_format, OF _ refl refl]: "\ \ (c, s) \\<^sup>* (c', t) \ \P x y cs. c = Parallel cs \ s = Normal x \ t = Normal y \ \,\\\<^bsub>/F \<^esub>P c Q,A \ final (c', t) \ x \ pre P \ t \ Fault ` F \ (c' = Throw \ t \ Normal ` A) \ (c' = Skip \ t \ Normal ` Q)" apply(erule converse_rtranclp_induct2) apply (clarsimp simp: final_def) apply(erule step.cases, simp_all) \ \Parallel\ apply clarsimp apply (frule Normal_pre_star) apply (drule oghoare_Parallel) apply clarsimp apply (rename_tac i cs c1' x y s' as) apply (subgoal_tac "\\ (Parallel cs, Normal x) \ (Parallel (cs[i := c1']), Normal s')") apply (frule_tac c="Parallel cs" and a="AnnPar as" and Q="(\x\set as. postcond x)" and A ="(\x\set as. abrcond x)" in oghoare_step[where \=\ and F=F]) apply(rule Parallel, simp) apply clarsimp apply (rule Conseq, rule exI[where x="{}"], fastforce) apply clarsimp apply force apply force apply clarsimp apply clarsimp apply (rename_tac a') apply (drule_tac x=a' in spec) apply (drule mp, rule Conseq) apply (rule_tac x="{}" in exI) apply (rule_tac x="(\x\set as. postcond x)" in exI) apply (rule_tac x="(\x\set as. abrcond x)" in exI) apply (simp) apply clarsimp apply(erule (1) step.Parallel) \ \ParSkip\ apply (frule no_steps_final, simp add: final_def) apply clarsimp apply (drule oghoare_Parallel) apply clarsimp apply (rule imageI) apply (erule subsetD) apply clarsimp apply (clarsimp simp: in_set_conv_nth) apply (rename_tac i) apply (frule_tac x="i" in spec) apply clarsimp apply (frule_tac x="cs!i" in bspec) apply (clarsimp simp: in_set_conv_nth) apply (rule_tac x="i" in exI) apply clarsimp apply clarsimp apply (drule_tac x="as ! i" in bspec) apply (clarsimp simp: in_set_conv_nth) apply fastforce apply (drule oghoare_Skip) apply fastforce \ \ParThrow\ apply clarsimp apply (frule no_steps_final, simp add: final_def) apply clarsimp apply (drule oghoare_Parallel) apply (clarsimp simp: in_set_conv_nth) apply (drule_tac x=i in spec) apply clarsimp apply (drule oghoare_Throw) apply clarsimp apply (rename_tac i as Q' A' P') apply (drule_tac x="as ! i" in bspec) apply (clarsimp simp: in_set_conv_nth) apply (rule_tac x=i in exI, fastforce) apply clarsimp apply (rule imageI) apply (erule_tac A="(\x\set as. abrcond x)" in subsetD) apply clarsimp apply (rule_tac x="as!i" in bexI) apply blast apply clarsimp done lemma oghoare_step_Fault[rule_format, OF _ refl refl]: "\\ cf \ cf' \ cf = (c, Normal x) \ cf' = (c', Fault f) \ x \ pre P \ \,\\\<^bsub>/F \<^esub>P c Q,A \ f \ F" apply (induct arbitrary: x c c' P Q A f rule: step.induct, simp_all) apply clarsimp apply (drule oghoare_Guard) apply clarsimp apply blast apply clarsimp apply (drule oghoare_Seq) apply clarsimp apply clarsimp apply (drule oghoare_Catch) apply clarsimp apply clarsimp apply (rename_tac i cs c' x P Q A f) apply (drule oghoare_Parallel) apply clarsimp apply (rename_tac i cs c' x Q A f as) apply (drule_tac x="i" in spec) apply clarsimp apply (drule meta_spec)+ apply (erule meta_impE, rule conjI, (rule refl)+)+ apply (drule_tac x="as!i" in bspec) apply (clarsimp simp: in_set_conv_nth) apply (rule_tac x="i" in exI, fastforce) apply (erule (1) meta_impE) apply (erule (2) meta_impE) apply clarsimp apply (drule rtranclp_conjD[THEN conjunct1]) apply (drule oghoare_Await) apply clarsimp apply (rename_tac b c c' x Q A f r P' Q' A') apply (drule (1) oghoare_atom_com_sound) apply (simp add: valid_def) apply (drule_tac x="Normal x" in spec) apply (drule_tac x="Fault f" in spec) apply (drule_tac x=Skip in spec) apply clarsimp apply (erule impE) apply (cut_tac f=f and c=c' in steps_Fault[where \=\]) apply fastforce apply (fastforce simp: final_def steps_Fault) done lemma oghoare_step_Stuck[rule_format, OF _ refl refl]: "\\ cf \ cf' \ cf = (c, Normal x) \ cf' = (c', Stuck) \ x \ pre P \ \,\\\<^bsub>/F \<^esub>P c Q,A \ P'" apply (induct arbitrary: x c c' P Q A rule: step.induct, simp_all) apply clarsimp apply (drule oghoare_Spec) apply force apply clarsimp apply (drule oghoare_Seq) apply clarsimp apply clarsimp apply (drule oghoare_Call) apply clarsimp apply clarsimp apply (drule oghoare_Catch) apply clarsimp apply clarsimp apply (drule oghoare_Parallel) apply clarsimp apply (rename_tac i cs c' x Q A as) apply (drule_tac x="i" in spec) apply clarsimp apply (drule meta_spec)+ apply (erule meta_impE, rule conjI, (rule refl)+)+ apply (drule_tac x="as!i" in bspec) apply (clarsimp simp: in_set_conv_nth) apply (rule_tac x="i" in exI, fastforce) apply (erule meta_impE[OF _ refl]) apply (erule (1) meta_impE) apply (erule (2) meta_impE) apply clarsimp apply (drule rtranclp_conjD[THEN conjunct1]) apply (drule oghoare_Await) apply clarsimp apply (rename_tac b c c' x Q A r P'' Q' A') apply (drule (1) oghoare_atom_com_sound) apply (simp add: valid_def) apply (drule_tac x="Normal x" in spec) apply (drule_tac x=Stuck in spec) apply (drule_tac x=Skip in spec) apply clarsimp apply (erule impE) apply (cut_tac c=c' in steps_Stuck[where \=\]) apply fastforce apply (fastforce simp: final_def steps_Fault) apply clarsimp apply (drule oghoare_Await) apply clarsimp done lemma oghoare_steps_Fault[rule_format, OF _ refl refl]: "\\ cf \\<^sup>* cf' \ cf = (c, Normal x) \ cf' = (c', Fault f) \ x \ pre P \ \,\\\<^bsub>/F \<^esub>P c Q,A \ f \ F" apply (induct arbitrary: x c c' f rule: rtranclp_induct) apply fastforce apply clarsimp apply (rename_tac b x c c' f) apply (case_tac b) apply clarsimp apply (drule (2) oghoare_steps) apply clarsimp apply (drule (3) oghoare_step_Fault) apply clarsimp apply (drule meta_spec)+ apply (erule meta_impE, (rule conjI, (rule refl)+))+ apply simp apply (drule step_Fault_prop ; simp) apply simp apply clarsimp apply (drule step_Stuck_prop ; simp) done lemma oghoare_steps_Stuck[rule_format, OF _ refl refl]: "\\ cf \\<^sup>* cf' \ cf = (c, Normal x) \ cf' = (c', Stuck) \ x \ pre P \ \,\\\<^bsub>/F \<^esub>P c Q,A \ P'" apply (induct arbitrary: x c c' rule: rtranclp_induct) apply fastforce apply clarsimp apply (rename_tac b x c c') apply (case_tac b) apply clarsimp apply (drule (2) oghoare_steps) apply clarsimp apply (drule (3) oghoare_step_Stuck) apply clarsimp - apply (drule meta_spec)+ - apply (erule meta_impE, (rule conjI, (rule refl)+))+ - apply simp apply (drule step_Fault_prop ; simp) apply simp done lemma oghoare_sound_Parallel_Fault_case[rule_format, OF _ refl refl]: "\ \ (c, s) \\<^sup>* (c', t) \ \P x f cs. c = Parallel cs \ s = Normal x \ x \ pre P \ t = Fault f \ \,\\\<^bsub>/F \<^esub>P c Q,A \ final (c', t) \ f \ F" apply(erule converse_rtranclp_induct2) apply (clarsimp simp: final_def) apply clarsimp apply (rename_tac c s P x f cs) apply (case_tac s) apply clarsimp apply(erule step.cases, simp_all) apply (clarsimp simp: final_def) apply (drule oghoare_Parallel) apply clarsimp apply (rename_tac x f s' i cs c1' as) apply (subgoal_tac "\\ (Parallel cs, Normal x) \ (Parallel (cs[i := c1']), Normal s')") apply (frule_tac c="Parallel cs" and a="AnnPar as" and Q="(\x\set as. postcond x)" and A="(\x\set as. abrcond x)" in oghoare_step[where \=\ and F=F]) apply(rule Parallel) apply simp apply clarsimp apply (rule Conseq, rule exI[where x="{}"], fastforce) apply assumption apply clarsimp apply clarsimp apply simp apply clarsimp apply (rename_tac a') apply (drule_tac x=a' in spec) apply clarsimp apply (erule notE[where P="oghoare _ _ _ _ _ _ _"]) apply (rule Conseq, rule exI[where x="{}"]) apply (clarsimp) apply (rule_tac x="(\x\set as. postcond x)" in exI) apply (rule_tac x="(\x\set as. abrcond x)" in exI ; simp) apply(erule (1) step.Parallel) apply clarsimp apply (fastforce dest: no_steps_final simp: final_def)+ apply (clarsimp simp: final_def) apply (drule oghoare_Parallel) apply (erule step_Normal_elim_cases, simp_all) apply clarsimp apply (rename_tac f cs f' i c1' as) apply (drule_tac x="i" in spec) apply (erule impE, fastforce) apply clarsimp apply (drule_tac x="as!i" in bspec) apply (clarsimp simp: in_set_conv_nth) apply (rule_tac x="i" in exI, fastforce) apply (drule_tac P="pres (as ! i)" in oghoare_step_Fault[where \=\ and F=F]) apply assumption+ apply (drule steps_Fault_prop ; simp) apply simp apply (drule steps_Stuck_prop ;simp) done lemma oghoare_soundness: "(\, \ \\<^bsub>/F\<^esub> P c Q,A \ \ \\<^bsub>/F\<^esub> (pre P) c Q, A) \ (\, \ \\<^bsub>/F\<^esub>P' P c Q,A \ \ \\<^bsub>/F\<^esub> P' c Q, A)" unfolding valid_def proof (induct rule: oghoare_oghoare_seq.induct) case SeqSkip thus ?case by (fastforce elim: converse_rtranclpE step_Normal_elim_cases(1)) next case SeqThrow thus ?case by (fastforce elim: converse_rtranclpE step_Normal_elim_cases) next case SeqBasic thus ?case by (fastforce elim: converse_rtranclpE step_Normal_elim_cases simp: final_def) next case (SeqSpec \ \ F r Q A) thus ?case apply clarsimp apply (erule converse_rtranclpE) apply (clarsimp simp: final_def) apply (erule step_Normal_elim_cases) apply (fastforce elim!: converse_rtranclpE step_Normal_elim_cases) by clarsimp next case (SeqSeq \ \ F P\<^sub>1 c\<^sub>1 P\<^sub>2 A c\<^sub>2 Q) show ?case using SeqSeq by (fold valid_def) (fastforce intro: oghoare_seq_valid simp: valid_weaken_pre) next case (SeqCatch \ \ F P\<^sub>1 c\<^sub>1 Q P\<^sub>2 c\<^sub>2 A) thus ?case by (fold valid_def) (fastforce elim: oghoare_catch_valid)+ next case (SeqCond \ \ F P b c1 Q A c2) thus ?case by (fold valid_def) (fastforce intro:oghoare_if_valid) next case (SeqWhile \ \ F P c A b) thus ?case by (fold valid_def) (fastforce elim: valid_weaken_pre[rotated] oghoare_while_valid) next case (SeqGuard \ \ F P c Q A r g f) thus ?case apply (fold valid_def) apply (simp (no_asm) add: valid_def) apply clarsimp apply (erule converse_rtranclpE) apply (fastforce simp: final_def) apply clarsimp apply (erule step_Normal_elim_cases) apply (case_tac "r \ - g \ {}") apply clarsimp apply (fastforce simp: valid_def) apply clarsimp apply (fastforce simp: valid_def) apply clarsimp apply (case_tac "r \ - g \ {}") apply (fastforce dest: no_steps_final simp:final_def) apply (fastforce dest: no_steps_final simp:final_def) done next case (SeqCall \ p f \ F P Q A) thus ?case apply clarsimp apply (erule converse_rtranclpE) apply (clarsimp simp add: final_def) apply (erule step_Normal_elim_cases) apply (clarsimp simp: final_def) apply fastforce apply fastforce done next case (SeqDynCom r P fa \ \ F c Q A) thus ?case apply - apply clarsimp apply (erule converse_rtranclpE) apply (clarsimp simp: final_def) apply clarsimp apply (erule step_Normal_elim_cases) apply clarsimp apply (rename_tac t c' x) apply (drule_tac x=x in bspec, fastforce) apply clarsimp apply (drule_tac x="Normal x" in spec) apply (drule_tac x="t" in spec) apply (drule_tac x="c'" in spec) apply fastforce+ done next case (SeqConseq \ \ F P c Q A) thus ?case apply (clarsimp) apply (rename_tac t c' x) apply (erule_tac x="Normal x" in allE) apply (erule_tac x="t" in allE) apply (erule_tac x="c'" in allE) apply (clarsimp simp: pre_weaken_pre) apply force done next case (SeqParallel as P \ \ F cs Q A) thus ?case by (fold valid_def) (erule (1) valid_weaken_pre) next case (Call \ p as n P Q A r \ f F) thus ?case apply clarsimp apply (erule converse_rtranclpE) apply (clarsimp simp add: final_def) apply (erule step_Normal_elim_cases) apply (clarsimp simp: final_def) apply (erule disjE) apply clarsimp apply fastforce apply fastforce apply fastforce done next case (Await \ \ F P c Q A r b) thus ?case apply clarsimp apply (erule converse_rtranclpE) apply (clarsimp simp add: final_def) apply (erule step_Normal_elim_cases) apply (erule converse_rtranclpE) apply (fastforce simp add: final_def ) apply (force dest!:no_step_final simp: final_def) apply clarsimp apply (rename_tac x c'') apply (drule_tac x="Normal x" in spec) apply (drule_tac x="Stuck" in spec) apply (drule_tac x="Skip" in spec) apply (clarsimp simp: final_def) apply (erule impE[where P="rtranclp _ _ _"]) apply (cut_tac c="c''" in steps_Stuck[where \=\]) apply fastforce apply fastforce apply clarsimp apply (rename_tac x c'' f) apply (drule_tac x="Normal x" in spec) apply (drule_tac x="Fault f" in spec) apply (drule_tac x="Skip" in spec) apply (erule impE[where P="rtranclp _ _ _"]) apply (cut_tac c="c''" and f=f in steps_Fault[where \=\]) apply fastforce apply clarsimp apply (erule converse_rtranclpE) apply fastforce apply (erule step_elim_cases) apply (fastforce) done next case (Parallel as cs \ \ F Q A ) thus ?case apply (fold valid_def) apply (simp only:pre.simps) apply (simp (no_asm) only: valid_def) apply clarsimp apply (rename_tac t c' x') apply (case_tac t) apply clarsimp apply (drule oghoare_sound_Parallel_Normal_case[where \=\ and Q=Q and A=A and F=F and P="AnnPar as", OF _ refl]) apply (rule oghoare_oghoare_seq.Parallel) apply simp apply clarsimp apply assumption apply (clarsimp) apply clarsimp apply (clarsimp simp: final_def) apply (clarsimp ) apply clarsimp apply clarsimp apply (drule oghoare_sound_Parallel_Fault_case[where \=\ and Q=Q and A=A and F=F and P="AnnPar as", OF _ ]) apply clarsimp apply assumption apply (rule oghoare_oghoare_seq.Parallel) apply simp apply clarsimp apply assumption apply clarsimp apply clarsimp apply (simp add: final_def) apply (fastforce simp add: final_def) apply (clarsimp simp: final_def) apply (erule oghoare_steps_Stuck[where \=\ and F=F and Q=Q and A=A and P="AnnPar as"]) apply simp apply (rule oghoare_oghoare_seq.Parallel) apply simp apply simp apply simp apply clarsimp apply clarsimp done next case Skip thus ?case by (fastforce elim: converse_rtranclpE step_Normal_elim_cases(1)) next case Basic thus ?case by (fastforce elim: converse_rtranclpE step_Normal_elim_cases simp: final_def) next case (Spec \ \ F r Q A) thus ?case apply clarsimp apply (erule converse_rtranclpE) apply (clarsimp simp: final_def) apply (erule step_Normal_elim_cases) apply (fastforce elim!: converse_rtranclpE step_Normal_elim_cases) by clarsimp next case (Seq \ \ F P\<^sub>1 c\<^sub>1 P\<^sub>2 A c\<^sub>2 Q) show ?case using Seq by (fold valid_def) (fastforce intro: oghoare_seq_valid simp: valid_weaken_pre) next case (Cond \ \ F P\<^sub>1 c\<^sub>1 Q A P\<^sub>2 c\<^sub>2 r b) thus ?case by (fold valid_def) (fastforce intro:oghoare_if_valid) next case (While \ \ F P c i A b Q r) thus ?case by (fold valid_def) (fastforce elim: valid_weaken_pre[rotated] oghoare_while_valid) next case Throw thus ?case by (fastforce elim: converse_rtranclpE step_Normal_elim_cases) next case (Catch \ \ F P\<^sub>1 c\<^sub>1 Q P\<^sub>2 c\<^sub>2 A) thus ?case apply (fold valid_def) apply (fastforce elim: oghoare_catch_valid)+ done next case (Guard \ \ F P c Q A r g f) thus ?case apply (fold valid_def) apply (simp) apply (frule (1) valid_weaken_pre[rotated]) apply (simp (no_asm) add: valid_def) apply clarsimp apply (erule converse_rtranclpE) apply (fastforce simp: final_def) apply clarsimp apply (erule step_Normal_elim_cases) apply (case_tac "r \ - g \ {}") apply clarsimp apply (fastforce simp: valid_def) apply clarsimp apply (fastforce simp: valid_def) apply clarsimp apply (case_tac "r \ - g \ {}") apply clarsimp apply (fastforce dest: no_steps_final simp:final_def) apply (clarsimp simp: ex_in_conv[symmetric]) done next case (DynCom r \ \ F P c Q A) thus ?case apply clarsimp apply (erule converse_rtranclpE) apply (clarsimp simp: final_def) apply clarsimp apply (erule step_Normal_elim_cases) apply clarsimp apply (rename_tac t c' x) apply (erule_tac x=x in ballE) apply clarsimp apply (drule_tac x="Normal x" in spec) apply (drule_tac x="t" in spec) apply (drule_tac x="c'" in spec) apply fastforce+ done next case (Conseq \ \ F P c Q A) thus ?case apply (clarsimp) apply (rename_tac P' Q' A' t c' x) apply (erule_tac x="Normal x" in allE) apply (erule_tac x="t" in allE) apply (erule_tac x="c'" in allE) apply (clarsimp simp: pre_weaken_pre) apply force done qed lemmas oghoare_sound = oghoare_soundness[THEN conjunct1, rule_format] lemmas oghoare_seq_sound = oghoare_soundness[THEN conjunct2, rule_format] end diff --git a/thys/IP_Addresses/WordInterval.thy b/thys/IP_Addresses/WordInterval.thy --- a/thys/IP_Addresses/WordInterval.thy +++ b/thys/IP_Addresses/WordInterval.thy @@ -1,776 +1,775 @@ (* Title: WordInterval.thy Authors: Julius Michaelis, Cornelius Diekmann *) theory WordInterval imports Main "Word_Lib.Word_Lemmas" begin section\WordInterval: Executable datatype for Machine Word Sets\ text\Stores ranges of machine words as interval. This has been proven quite efficient for IP Addresses.\ (*NOTE: All algorithms here use a straight-forward implementation. There is a lot of room for improving the computation complexity, for example by making the WordInterval a balanced, sorted tree.*) subsection\Syntax\ context notes [[typedef_overloaded]] begin datatype ('a::len0) wordinterval = WordInterval "('a::len0) word" \ \start (inclusive)\ "('a::len0) word" \ \end (inclusive)\ | RangeUnion "'a wordinterval" "'a wordinterval" end subsection\Semantics\ fun wordinterval_to_set :: "'a::len0 wordinterval \ ('a::len0 word) set" where "wordinterval_to_set (WordInterval start end) = {start .. end}" | "wordinterval_to_set (RangeUnion r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" (*Note: The runtime of all the operations could be improved, for example by keeping the tree sorted and balanced.*) subsection\Basic operations\ text\\\\\ fun wordinterval_element :: "'a::len0 word \ 'a::len0 wordinterval \ bool" where "wordinterval_element el (WordInterval s e) \ s \ el \ el \ e" | "wordinterval_element el (RangeUnion r1 r2) \ wordinterval_element el r1 \ wordinterval_element el r2" lemma wordinterval_element_set_eq[simp]: "wordinterval_element el rg = (el \ wordinterval_to_set rg)" by(induction rg rule: wordinterval_element.induct) simp_all definition wordinterval_union :: "'a::len0 wordinterval \ 'a::len0 wordinterval \ 'a::len0 wordinterval" where "wordinterval_union r1 r2 = RangeUnion r1 r2" lemma wordinterval_union_set_eq[simp]: "wordinterval_to_set (wordinterval_union r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" unfolding wordinterval_union_def by simp fun wordinterval_empty :: "'a::len0 wordinterval \ bool" where "wordinterval_empty (WordInterval s e) \ e < s" | "wordinterval_empty (RangeUnion r1 r2) \ wordinterval_empty r1 \ wordinterval_empty r2" lemma wordinterval_empty_set_eq[simp]: "wordinterval_empty r \ wordinterval_to_set r = {}" by(induction r) auto definition Empty_WordInterval :: "'a::len wordinterval" where "Empty_WordInterval \ WordInterval 1 0" lemma wordinterval_empty_Empty_WordInterval: "wordinterval_empty Empty_WordInterval" by(simp add: Empty_WordInterval_def) lemma Empty_WordInterval_set_eq[simp]: "wordinterval_to_set Empty_WordInterval = {}" by(simp add: Empty_WordInterval_def) subsection\WordInterval and Lists\ text\A list of \(start, end)\ tuples.\ text\wordinterval to list\ fun wi2l :: "'a::len0 wordinterval \ ('a::len0 word \ 'a::len0 word) list" where "wi2l (RangeUnion r1 r2) = wi2l r1 @ wi2l r2" | "wi2l (WordInterval s e) = (if e < s then [] else [(s,e)])" text\list to wordinterval\ fun l2wi :: "('a::len word \ 'a word) list \ 'a wordinterval" where "l2wi [] = Empty_WordInterval" | "l2wi [(s,e)] = (WordInterval s e)" | "l2wi ((s,e)#rs) = (RangeUnion (WordInterval s e) (l2wi rs))" lemma l2wi_append: "wordinterval_to_set (l2wi (l1@l2)) = wordinterval_to_set (l2wi l1) \ wordinterval_to_set (l2wi l2)" proof(induction l1 arbitrary: l2 rule:l2wi.induct) case 1 thus ?case by simp next case (2 s e l2) thus ?case by (cases l2) simp_all next case 3 thus ?case by force qed lemma l2wi_wi2l[simp]: "wordinterval_to_set (l2wi (wi2l r)) = wordinterval_to_set r" by(induction r) (simp_all add: l2wi_append) lemma l2wi: "wordinterval_to_set (l2wi l) = (\ (i,j) \ set l. {i .. j})" by(induction l rule: l2wi.induct, simp_all) lemma wi2l: "(\(i,j)\set (wi2l r). {i .. j}) = wordinterval_to_set r" by(induction r rule: wi2l.induct, simp_all) lemma l2wi_remdups[simp]: "wordinterval_to_set (l2wi (remdups ls)) = wordinterval_to_set (l2wi ls)" by(simp add: l2wi) lemma wi2l_empty[simp]: "wi2l Empty_WordInterval = []" unfolding Empty_WordInterval_def by simp subsection\Optimizing and minimizing @{typ "('a::len) wordinterval"}s\ text\Removing empty intervals\ context begin fun wordinterval_optimize_empty :: "'a::len0 wordinterval \ 'a wordinterval" where "wordinterval_optimize_empty (RangeUnion r1 r2) = (let r1o = wordinterval_optimize_empty r1; r2o = wordinterval_optimize_empty r2 in if wordinterval_empty r1o then r2o else if wordinterval_empty r2o then r1o else RangeUnion r1o r2o)" | "wordinterval_optimize_empty r = r" lemma wordinterval_optimize_empty_set_eq[simp]: "wordinterval_to_set (wordinterval_optimize_empty r) = wordinterval_to_set r" by(induction r) (simp_all add: Let_def) lemma wordinterval_optimize_empty_double: "wordinterval_optimize_empty (wordinterval_optimize_empty r) = wordinterval_optimize_empty r" by(induction r) (simp_all add: Let_def) private fun wordinterval_empty_shallow :: "'a::len0 wordinterval \ bool" where "wordinterval_empty_shallow (WordInterval s e) \ e < s" | "wordinterval_empty_shallow (RangeUnion _ _) \ False" private lemma helper_optimize_shallow: "wordinterval_empty_shallow (wordinterval_optimize_empty r) = wordinterval_empty (wordinterval_optimize_empty r)" by(induction r) fastforce+ private fun wordinterval_optimize_empty2 where "wordinterval_optimize_empty2 (RangeUnion r1 r2) = (let r1o = wordinterval_optimize_empty r1; r2o = wordinterval_optimize_empty r2 in if wordinterval_empty_shallow r1o then r2o else if wordinterval_empty_shallow r2o then r1o else RangeUnion r1o r2o)" | "wordinterval_optimize_empty2 r = r" lemma wordinterval_optimize_empty_code[code_unfold]: "wordinterval_optimize_empty = wordinterval_optimize_empty2" by (subst fun_eq_iff, clarify, rename_tac r, induct_tac r) (unfold wordinterval_optimize_empty.simps wordinterval_optimize_empty2.simps Let_def helper_optimize_shallow, simp_all) end text\Merging overlapping intervals\ context begin private definition disjoint :: "'a set \ 'a set \ bool" where "disjoint A B \ A \ B = {}" private primrec interval_of :: "('a::len0) word \ 'a word \ 'a word set" where "interval_of (s,e) = {s .. e}" declare interval_of.simps[simp del] private definition disjoint_intervals :: "(('a::len0) word \ ('a::len0) word) \ ('a word \ 'a word) \ bool" where "disjoint_intervals A B \ disjoint (interval_of A) (interval_of B)" private definition not_disjoint_intervals :: "(('a::len0) word \ ('a::len0) word) \ ('a word \ 'a word) \ bool" where "not_disjoint_intervals A B \ \ disjoint (interval_of A) (interval_of B)" private lemma [code]: "not_disjoint_intervals A B = (case A of (s,e) \ case B of (s',e') \ s \ e' \ s' \ e \ s \ e \ s' \ e')" apply(cases A, cases B) apply(simp add: not_disjoint_intervals_def interval_of.simps disjoint_def) done private lemma [code]: "disjoint_intervals A B = (case A of (s,e) \ case B of (s',e') \ s > e' \ s' > e \ s > e \ s' > e')" apply(cases A, cases B) apply(simp add: disjoint_intervals_def interval_of.simps disjoint_def) by fastforce text\BEGIN merging overlapping intervals\ (*result has no empty intervals and all are disjoint. merging things such as [1,7] [8,10] would still be possible*) private fun merge_overlap :: "(('a::len0) word \ ('a::len0) word) \ ('a word \ 'a word) list \ ('a word \ 'a word) list" where "merge_overlap s [] = [s]" | "merge_overlap (s,e) ((s',e')#ss) = ( if not_disjoint_intervals (s,e) (s',e') then (min s s', max e e')#ss else (s',e')#merge_overlap (s,e) ss)" private lemma not_disjoint_union: fixes s :: "('a::len0) word" shows "\ disjoint {s..e} {s'..e'} \ {s..e} \ {s'..e'} = {min s s' .. max e e'}" by(auto simp add: disjoint_def min_def max_def) private lemma disjoint_subset: "disjoint A B \ A \ B \ C \ A \ C" unfolding disjoint_def by blast private lemma merge_overlap_helper1: "interval_of A \ (\s \ set ss. interval_of s) \ (\s \ set (merge_overlap A ss). interval_of s) = (\s \ set ss. interval_of s)" apply(induction ss) apply(simp; fail) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: not_disjoint_intervals_def interval_of.simps) apply(intro impI conjI) apply(drule not_disjoint_union) apply blast apply(drule_tac C="(\x\set xs. interval_of x)" in disjoint_subset) apply(simp_all) done private lemma merge_overlap_helper2: "\s'\set ss. \ disjoint (interval_of A) (interval_of s') \ interval_of A \ (\s \ set ss. interval_of s) = (\s \ set (merge_overlap A ss). interval_of s)" apply(induction ss) apply(simp; fail) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: not_disjoint_intervals_def interval_of.simps) apply(intro impI conjI) apply(drule not_disjoint_union) apply blast apply(simp) by blast private lemma merge_overlap_length: "\s' \ set ss. \ disjoint (interval_of A) (interval_of s') \ length (merge_overlap A ss) = length ss" apply(induction ss) apply(simp) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: not_disjoint_intervals_def interval_of.simps) done lemma "merge_overlap (1:: 16 word,2) [(1, 7)] = [(1, 7)]" by eval lemma "merge_overlap (1:: 16 word,2) [(2, 7)] = [(1, 7)]" by eval lemma "merge_overlap (1:: 16 word,2) [(3, 7)] = [(3, 7), (1,2)]" by eval private function listwordinterval_compress :: "(('a::len0) word \ ('a::len0) word) list \ ('a word \ 'a word) list" where "listwordinterval_compress [] = []" | "listwordinterval_compress (s#ss) = ( if \s' \ set ss. disjoint_intervals s s' then s#listwordinterval_compress ss else listwordinterval_compress (merge_overlap s ss))" by(pat_completeness, auto) private termination listwordinterval_compress apply (relation "measure length") apply(rule wf_measure) apply(simp) using disjoint_intervals_def merge_overlap_length by fastforce private lemma listwordinterval_compress: "(\s \ set (listwordinterval_compress ss). interval_of s) = (\s \ set ss. interval_of s)" apply(induction ss rule: listwordinterval_compress.induct) apply(simp) apply(simp) apply(intro impI) apply(simp add: disjoint_intervals_def) apply(drule merge_overlap_helper2) apply(simp) done lemma "listwordinterval_compress [(1::32 word,3), (8,10), (2,5), (3,7)] = [(8, 10), (1, 7)]" by eval private lemma A_in_listwordinterval_compress: "A \ set (listwordinterval_compress ss) \ interval_of A \ (\s \ set ss. interval_of s)" using listwordinterval_compress by blast private lemma listwordinterval_compress_disjoint: "A \ set (listwordinterval_compress ss) \ B \ set (listwordinterval_compress ss) \ A \ B \ disjoint (interval_of A) (interval_of B)" apply(induction ss arbitrary: rule: listwordinterval_compress.induct) apply(simp) apply(simp split: if_split_asm) apply(elim disjE) apply(simp_all) apply(simp_all add: disjoint_intervals_def disjoint_def) - apply(thin_tac [!] "False \ _ \ _ \ _") apply(blast dest: A_in_listwordinterval_compress)+ done text\END merging overlapping intervals\ text\BEGIN merging adjacent intervals\ private fun merge_adjacent :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) list \ ('a word \ 'a word) list" where "merge_adjacent s [] = [s]" | "merge_adjacent (s,e) ((s',e')#ss) = ( if s \e \ s' \ e' \ word_next e = s' then (s, e')#ss else if s \e \ s' \ e' \ word_next e' = s then (s', e)#ss else (s',e')#merge_adjacent (s,e) ss)" private lemma merge_adjacent_helper: "interval_of A \ (\s \ set ss. interval_of s) = (\s \ set (merge_adjacent A ss). interval_of s)" apply(induction ss) apply(simp; fail) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: interval_of.simps) apply(intro impI conjI) apply (metis Un_assoc word_adjacent_union) apply(elim conjE) apply(drule(2) word_adjacent_union) subgoal by (blast) subgoal by (metis word_adjacent_union Un_assoc) by blast private lemma merge_adjacent_length: "\(s', e')\set ss. s \ e \ s' \ e' \ (word_next e = s' \ word_next e' = s) \ length (merge_adjacent (s,e) ss) = length ss" apply(induction ss) apply(simp) apply(rename_tac x xs) apply(case_tac x) apply(simp add: ) by blast private function listwordinterval_adjacent :: "(('a::len) word \ ('a::len) word) list \ ('a word \ 'a word) list" where "listwordinterval_adjacent [] = []" | "listwordinterval_adjacent ((s,e)#ss) = ( if \(s',e') \ set ss. \ (s \e \ s' \ e' \ (word_next e = s' \ word_next e' = s)) then (s,e)#listwordinterval_adjacent ss else listwordinterval_adjacent (merge_adjacent (s,e) ss))" by(pat_completeness, auto) private termination listwordinterval_adjacent apply (relation "measure length") apply(rule wf_measure) apply(simp) apply(simp) using merge_adjacent_length by fastforce private lemma listwordinterval_adjacent: "(\s \ set (listwordinterval_adjacent ss). interval_of s) = (\s \ set ss. interval_of s)" apply(induction ss rule: listwordinterval_adjacent.induct) apply(simp) apply(simp add: merge_adjacent_helper) done lemma "listwordinterval_adjacent [(1::16 word, 3), (5, 10), (10,10), (4,4)] = [(10, 10), (1, 10)]" by eval text\END merging adjacent intervals\ definition wordinterval_compress :: "('a::len) wordinterval \ 'a wordinterval" where "wordinterval_compress r \ l2wi (remdups (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))))" text\Correctness: Compression preserves semantics\ lemma wordinterval_compress: "wordinterval_to_set (wordinterval_compress r) = wordinterval_to_set r" unfolding wordinterval_compress_def proof - have interval_of': "interval_of s = (case s of (s,e) \ {s .. e})" for s by (cases s) (simp add: interval_of.simps) have "wordinterval_to_set (l2wi (remdups (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))))) = (\x\set (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))). interval_of x)" by (force simp: interval_of' l2wi) also have "\ = (\s\set (wi2l (wordinterval_optimize_empty r)). interval_of s)" by(simp add: listwordinterval_compress listwordinterval_adjacent) also have "\ = (\(i, j)\set (wi2l (wordinterval_optimize_empty r)). {i..j})" by(simp add: interval_of') also have "\ = wordinterval_to_set r" by(simp add: wi2l) finally show "wordinterval_to_set (l2wi (remdups (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))))) = wordinterval_to_set r" . qed end text\Example\ lemma "(wi2l \ (wordinterval_compress :: 32 wordinterval \ 32 wordinterval) \ l2wi) [(70, 80001), (0,0), (150, 8000), (1,3), (42,41), (3,7), (56, 200), (8,10)] = [(56, 80001), (0, 10)]" by eval lemma "wordinterval_compress (RangeUnion (RangeUnion (WordInterval (1::32 word) 5) (WordInterval 8 10)) (WordInterval 3 7)) = WordInterval 1 10" by eval subsection\Further operations\ text\\\\\ definition wordinterval_Union :: "('a::len) wordinterval list \ 'a wordinterval" where "wordinterval_Union ws = wordinterval_compress (foldr wordinterval_union ws Empty_WordInterval)" lemma wordinterval_Union: "wordinterval_to_set (wordinterval_Union ws) = (\ w \ (set ws). wordinterval_to_set w)" by(induction ws) (simp_all add: wordinterval_compress wordinterval_Union_def) context begin private fun wordinterval_setminus' :: "'a::len wordinterval \ 'a wordinterval \ 'a wordinterval" where "wordinterval_setminus' (WordInterval s e) (WordInterval ms me) = ( if s > e \ ms > me then WordInterval s e else if me \ e then WordInterval (if ms = 0 then 1 else s) (min e (word_prev ms)) else if ms \ s then WordInterval (max s (word_next me)) (if me = max_word then 0 else e) else RangeUnion (WordInterval (if ms = 0 then 1 else s) (word_prev ms)) (WordInterval (word_next me) (if me = max_word then 0 else e)) )" | "wordinterval_setminus' (RangeUnion r1 r2) t = RangeUnion (wordinterval_setminus' r1 t) (wordinterval_setminus' r2 t)"| "wordinterval_setminus' t (RangeUnion r1 r2) = wordinterval_setminus' (wordinterval_setminus' t r1) r2" private lemma wordinterval_setminus'_rr_set_eq: "wordinterval_to_set(wordinterval_setminus' (WordInterval s e) (WordInterval ms me)) = wordinterval_to_set (WordInterval s e) - wordinterval_to_set (WordInterval ms me)" apply(simp only: wordinterval_setminus'.simps) apply(case_tac "e < s") apply simp apply(case_tac "me < ms") apply simp apply(case_tac [!] "e \ me") apply(case_tac [!] "ms = 0") apply(case_tac [!] "ms \ s") apply(case_tac [!] "me = max_word") apply(simp_all add: word_prev_def word_next_def min_def max_def) apply(safe) apply(auto) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) done private lemma wordinterval_setminus'_set_eq: "wordinterval_to_set (wordinterval_setminus' r1 r2) = wordinterval_to_set r1 - wordinterval_to_set r2" apply(induction rule: wordinterval_setminus'.induct) using wordinterval_setminus'_rr_set_eq apply blast apply auto done lemma wordinterval_setminus'_empty_struct: "wordinterval_empty r2 \ wordinterval_setminus' r1 r2 = r1" by(induction r1 r2 rule: wordinterval_setminus'.induct) auto definition wordinterval_setminus :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_setminus r1 r2 = wordinterval_compress (wordinterval_setminus' r1 r2)" lemma wordinterval_setminus_set_eq[simp]: "wordinterval_to_set (wordinterval_setminus r1 r2) = wordinterval_to_set r1 - wordinterval_to_set r2" by(simp add: wordinterval_setminus_def wordinterval_compress wordinterval_setminus'_set_eq) end definition wordinterval_UNIV :: "'a::len wordinterval" where "wordinterval_UNIV \ WordInterval 0 max_word" lemma wordinterval_UNIV_set_eq[simp]: "wordinterval_to_set wordinterval_UNIV = UNIV" unfolding wordinterval_UNIV_def using max_word_max by fastforce fun wordinterval_invert :: "'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_invert r = wordinterval_setminus wordinterval_UNIV r" lemma wordinterval_invert_set_eq[simp]: "wordinterval_to_set (wordinterval_invert r) = UNIV - wordinterval_to_set r" by(auto) lemma wordinterval_invert_UNIV_empty: "wordinterval_empty (wordinterval_invert wordinterval_UNIV)" by simp lemma wi2l_univ[simp]: "wi2l wordinterval_UNIV = [(0, max_word)]" unfolding wordinterval_UNIV_def by simp text\\\\\ context begin private lemma "{(s::nat) .. e} \ {s' .. e'} = {} \ s > e' \ s' > e \ s > e \ s' > e'" by simp linarith private fun wordinterval_intersection' :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_intersection' (WordInterval s e) (WordInterval s' e') = ( if s > e \ s' > e' \ s > e' \ s' > e \ s > e \ s' > e' then Empty_WordInterval else WordInterval (max s s') (min e e') )" | "wordinterval_intersection' (RangeUnion r1 r2) t = RangeUnion (wordinterval_intersection' r1 t) (wordinterval_intersection' r2 t)"| "wordinterval_intersection' t (RangeUnion r1 r2) = RangeUnion (wordinterval_intersection' t r1) (wordinterval_intersection' t r2)" private lemma wordinterval_intersection'_set_eq: "wordinterval_to_set (wordinterval_intersection' r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" by(induction r1 r2 rule: wordinterval_intersection'.induct) (auto) lemma "wordinterval_intersection' (RangeUnion (RangeUnion (WordInterval (1::32 word) 3) (WordInterval 8 10)) (WordInterval 1 3)) (WordInterval 1 3) = RangeUnion (RangeUnion (WordInterval 1 3) (WordInterval 1 0)) (WordInterval 1 3)" by eval definition wordinterval_intersection :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_intersection r1 r2 \ wordinterval_compress (wordinterval_intersection' r1 r2)" lemma wordinterval_intersection_set_eq[simp]: "wordinterval_to_set (wordinterval_intersection r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" by(simp add: wordinterval_intersection_def wordinterval_compress wordinterval_intersection'_set_eq) lemma "wordinterval_intersection (RangeUnion (RangeUnion (WordInterval (1::32 word) 3) (WordInterval 8 10)) (WordInterval 1 3)) (WordInterval 1 3) = WordInterval 1 3" by eval end definition wordinterval_subset :: "'a::len wordinterval \ 'a::len wordinterval \ bool" where "wordinterval_subset r1 r2 \ wordinterval_empty (wordinterval_setminus r1 r2)" lemma wordinterval_subset_set_eq[simp]: "wordinterval_subset r1 r2 = (wordinterval_to_set r1 \ wordinterval_to_set r2)" unfolding wordinterval_subset_def by simp definition wordinterval_eq :: "'a::len wordinterval \ 'a::len wordinterval \ bool" where "wordinterval_eq r1 r2 = (wordinterval_subset r1 r2 \ wordinterval_subset r2 r1)" lemma wordinterval_eq_set_eq: "wordinterval_eq r1 r2 \ wordinterval_to_set r1 = wordinterval_to_set r2" unfolding wordinterval_eq_def by auto thm iffD1[OF wordinterval_eq_set_eq] (*declare iffD1[OF wordinterval_eq_set_eq, simp]*) lemma wordinterval_eq_comm: "wordinterval_eq r1 r2 \ wordinterval_eq r2 r1" unfolding wordinterval_eq_def by fast lemma wordinterval_to_set_alt: "wordinterval_to_set r = {x. wordinterval_element x r}" unfolding wordinterval_element_set_eq by blast lemma wordinterval_un_empty: "wordinterval_empty r1 \ wordinterval_eq (wordinterval_union r1 r2) r2" by(subst wordinterval_eq_set_eq, simp) lemma wordinterval_un_emty_b: "wordinterval_empty r2 \ wordinterval_eq (wordinterval_union r1 r2) r1" by(subst wordinterval_eq_set_eq, simp) lemma wordinterval_Diff_triv: "wordinterval_empty (wordinterval_intersection a b) \ wordinterval_eq (wordinterval_setminus a b) a" unfolding wordinterval_eq_set_eq by simp blast text\A size of the datatype, does not correspond to the cardinality of the corresponding set\ fun wordinterval_size :: "('a::len) wordinterval \ nat" where "wordinterval_size (RangeUnion a b) = wordinterval_size a + wordinterval_size b" | "wordinterval_size (WordInterval s e) = (if s \ e then 1 else 0)" lemma wordinterval_size_length: "wordinterval_size r = length (wi2l r)" by(induction r) (auto) lemma Ex_wordinterval_nonempty: "\x::('a::len wordinterval). y \ wordinterval_to_set x" proof show "y \ wordinterval_to_set wordinterval_UNIV" by simp qed lemma wordinterval_eq_reflp: "reflp wordinterval_eq" apply(rule reflpI) by(simp only: wordinterval_eq_set_eq) lemma wordintervalt_eq_symp: "symp wordinterval_eq" apply(rule sympI) by(simp add: wordinterval_eq_comm) lemma wordinterval_eq_transp: "transp wordinterval_eq" apply(rule transpI) by(simp only: wordinterval_eq_set_eq) lemma wordinterval_eq_equivp: "equivp wordinterval_eq" by (auto intro: equivpI wordinterval_eq_reflp wordintervalt_eq_symp wordinterval_eq_transp) text\The smallest element in the interval\ definition is_lowest_element :: "'a::ord \ 'a set \ bool" where "is_lowest_element x S = (x \ S \ (\y\S. y \ x \ y = x))" lemma fixes x :: "'a :: complete_lattice" assumes "x \ S" shows " x = Inf S \ is_lowest_element x S" using assms apply(simp add: is_lowest_element_def) by (simp add: Inf_lower eq_iff) lemma fixes x :: "'a :: linorder" assumes "finite S" and "x \ S" shows "is_lowest_element x S \ x = Min S" apply(rule) subgoal apply(simp add: is_lowest_element_def) apply(subst Min_eqI[symmetric]) using assms by(auto) by (metis Min.coboundedI assms(1) assms(2) dual_order.antisym is_lowest_element_def) text\Smallest element in the interval\ fun wordinterval_lowest_element :: "'a::len0 wordinterval \ 'a word option" where "wordinterval_lowest_element (WordInterval s e) = (if s \ e then Some s else None)" | "wordinterval_lowest_element (RangeUnion A B) = (case (wordinterval_lowest_element A, wordinterval_lowest_element B) of (Some a, Some b) \ Some (if a < b then a else b) | (None, Some b) \ Some b | (Some a, None) \ Some a | (None, None) \ None)" lemma wordinterval_lowest_none_empty: "wordinterval_lowest_element r = None \ wordinterval_empty r" proof(induction r) case WordInterval thus ?case by simp next case RangeUnion thus ?case by fastforce qed lemma wordinterval_lowest_element_correct_A: "wordinterval_lowest_element r = Some x \ is_lowest_element x (wordinterval_to_set r)" unfolding is_lowest_element_def apply(induction r arbitrary: x rule: wordinterval_lowest_element.induct) apply(rename_tac rs re x, case_tac "rs \ re", auto)[1] apply(subst(asm) wordinterval_lowest_element.simps(2)) apply(rename_tac A B x) apply(case_tac "wordinterval_lowest_element B") apply(case_tac[!] "wordinterval_lowest_element A") apply(simp_all add: wordinterval_lowest_none_empty)[3] apply fastforce done lemma wordinterval_lowest_element_set_eq: assumes "\ wordinterval_empty r" shows "(wordinterval_lowest_element r = Some x) = (is_lowest_element x (wordinterval_to_set r))" (*unfolding is_lowest_element_def*) proof(rule iffI) assume "wordinterval_lowest_element r = Some x" thus "is_lowest_element x (wordinterval_to_set r)" using wordinterval_lowest_element_correct_A wordinterval_lowest_none_empty by simp next assume "is_lowest_element x (wordinterval_to_set r)" with assms show "(wordinterval_lowest_element r = Some x)" proof(induction r arbitrary: x rule: wordinterval_lowest_element.induct) case 1 thus ?case by(simp add: is_lowest_element_def) next case (2 A B x) have is_lowest_RangeUnion: "is_lowest_element x (wordinterval_to_set A \ wordinterval_to_set B) \ is_lowest_element x (wordinterval_to_set A) \ is_lowest_element x (wordinterval_to_set B)" by(simp add: is_lowest_element_def) (*why \ A B?*) have wordinterval_lowest_element_RangeUnion: "\a b A B. wordinterval_lowest_element A = Some a \ wordinterval_lowest_element B = Some b \ wordinterval_lowest_element (RangeUnion A B) = Some (min a b)" by(auto dest!: wordinterval_lowest_element_correct_A simp add: is_lowest_element_def min_def) from 2 show ?case apply(case_tac "wordinterval_lowest_element B") apply(case_tac[!] "wordinterval_lowest_element A") apply(auto simp add: is_lowest_element_def)[3] apply(subgoal_tac "\ wordinterval_empty A \ \ wordinterval_empty B") prefer 2 using arg_cong[where f = Not, OF wordinterval_lowest_none_empty] apply blast apply(drule(1) wordinterval_lowest_element_RangeUnion) apply(simp split: option.split_asm add: min_def) apply(drule is_lowest_RangeUnion) apply(elim disjE) apply(simp add: is_lowest_element_def) apply(clarsimp simp add: wordinterval_lowest_none_empty) apply(simp add: is_lowest_element_def) apply(clarsimp simp add: wordinterval_lowest_none_empty) using wordinterval_lowest_element_correct_A[simplified is_lowest_element_def] by (metis Un_iff not_le) qed qed text\Cardinality approximation for @{typ "('a::len) wordinterval"}s\ context begin lemma card_atLeastAtMost_word: fixes s::"('a::len) word" shows "card {s..e} = Suc (unat e) - (unat s)" apply(cases "s > e") apply(simp) apply(subst(asm) Word.word_less_nat_alt) apply simp apply(subst Word_Lemmas.upto_enum_set_conv2[symmetric]) apply(subst List.card_set) apply(simp add: remdups_enum_upto) done fun wordinterval_card :: "('a::len) wordinterval \ nat" where "wordinterval_card (WordInterval s e) = Suc (unat e) - (unat s)" | "wordinterval_card (RangeUnion a b) = wordinterval_card a + wordinterval_card b" lemma wordinterval_card: "wordinterval_card r \ card (wordinterval_to_set r)" proof(induction r) case WordInterval thus ?case by (simp add: card_atLeastAtMost_word) next case (RangeUnion r1 r2) have "card (wordinterval_to_set r1 \ wordinterval_to_set r2) \ card (wordinterval_to_set r1) + card (wordinterval_to_set r2)" using Finite_Set.card_Un_le by blast with RangeUnion show ?case by(simp) qed text\With @{thm wordinterval_compress} it should be possible to get the exact cardinality\ end end diff --git a/thys/Iptables_Semantics/Semantics_Embeddings.thy b/thys/Iptables_Semantics/Semantics_Embeddings.thy --- a/thys/Iptables_Semantics/Semantics_Embeddings.thy +++ b/thys/Iptables_Semantics/Semantics_Embeddings.thy @@ -1,348 +1,347 @@ theory Semantics_Embeddings imports "Simple_Firewall/SimpleFw_Compliance" Matching_Embeddings Semantics "Semantics_Ternary/Semantics_Ternary" begin section\Semantics Embedding\ subsection\Tactic @{const in_doubt_allow}\ lemma iptables_bigstep_undecided_to_undecided_in_doubt_allow_approx: assumes agree: "matcher_agree_on_exact_matches \ \" and good: "good_ruleset rs" and semantics: "\,\,p\ \rs, Undecided\ \ Undecided" shows "(\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Undecided \ (\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow" proof - from semantics good show ?thesis proof(induction rs Undecided Undecided rule: iptables_bigstep_induct) case Skip thus ?case by(auto intro: approximating_bigstep.skip) next case Log thus ?case by(auto intro: approximating_bigstep.empty approximating_bigstep.log approximating_bigstep.nomatch) next case (Nomatch m a) with not_exact_match_in_doubt_allow_approx_match[OF agree] have "a \ Log \ a \ Empty \ a = Accept \ Matching_Ternary.matches (\, in_doubt_allow) m a p \ \ Matching_Ternary.matches (\, in_doubt_allow) m a p" by(simp add: good_ruleset_alt) blast thus ?case by(cases a) (auto intro: approximating_bigstep.empty approximating_bigstep.log approximating_bigstep.accept approximating_bigstep.nomatch) next case (Seq rs rs1 rs2 t) from Seq have "good_ruleset rs1" and "good_ruleset rs2" by(simp_all add: good_ruleset_append) also from Seq iptables_bigstep_to_undecided have "t = Undecided" by simp ultimately show ?case using Seq by(fastforce intro: approximating_bigstep.decision Semantics_Ternary.seq') qed(simp_all add: good_ruleset_def) qed lemma FinalAllow_approximating_in_doubt_allow: assumes agree: "matcher_agree_on_exact_matches \ \" and good: "good_ruleset rs" and semantics: "\,\,p\ \rs, Undecided\ \ Decision FinalAllow" shows "(\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow" proof - from semantics good show ?thesis proof(induction rs Undecided "Decision FinalAllow" rule: iptables_bigstep_induct) case Allow thus ?case by (auto intro: agree approximating_bigstep.accept in_doubt_allow_allows_Accept) next case (Seq rs rs1 rs2 t) from Seq have good1: "good_ruleset rs1" and good2: "good_ruleset rs2" by(simp_all add: good_ruleset_append) show ?case proof(cases t) case Decision with Seq good1 good2 show "(\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow" by (auto intro: approximating_bigstep.decision approximating_bigstep.seq dest: Semantics.decisionD) next case Undecided with iptables_bigstep_undecided_to_undecided_in_doubt_allow_approx[OF agree good1] Seq have "(\, in_doubt_allow),p\ \rs1, Undecided\ \\<^sub>\ Undecided \ (\, in_doubt_allow),p\ \rs1, Undecided\ \\<^sub>\ Decision FinalAllow" by simp with Undecided Seq good1 good2 show "(\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow" by (auto intro: approximating_bigstep.seq Semantics_Ternary.seq' approximating_bigstep.decision) qed next case Call_result thus ?case by(simp add: good_ruleset_alt) qed qed corollary FinalAllows_subseteq_in_doubt_allow: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ {p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow} \ {p. (\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow}" using FinalAllow_approximating_in_doubt_allow by (metis (lifting, full_types) Collect_mono) (*referenced by name in paper*) corollary new_packets_to_simple_firewall_overapproximation: defines "preprocess rs \ upper_closure (optimize_matches abstract_for_simple_firewall (upper_closure (packet_assume_new rs)))" and "newpkt p \ match_tcp_flags ipt_tcp_syn (p_tcp_flags p) \ p_tag_ctstate p = CT_New" fixes p :: "('i::len, 'pkt_ext) tagged_packet_scheme" assumes "matcher_agree_on_exact_matches \ common_matcher" and "simple_ruleset rs" shows "{p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow \ newpkt p} \ {p. simple_fw (to_simple_firewall (preprocess rs)) p = Decision FinalAllow \ newpkt p}" proof - from assms(3) have "{p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow \ newpkt p} \ {p. (common_matcher, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p}" apply(drule_tac rs=rs and \=\ in FinalAllows_subseteq_in_doubt_allow) using simple_imp_good_ruleset assms(4) apply blast by blast thus ?thesis unfolding newpkt_def preprocess_def using transform_simple_fw_upper(2)[OF assms(4)] by blast qed lemma approximating_bigstep_undecided_to_undecided_in_doubt_allow_approx: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ (\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Undecided \ \,\,p\ \rs, Undecided\ \ Undecided \ \,\,p\ \rs, Undecided\ \ Decision FinalDeny" apply(rotate_tac 2) apply(induction rs Undecided Undecided rule: approximating_bigstep_induct) apply(simp_all) apply (metis iptables_bigstep.skip) apply (metis iptables_bigstep.empty iptables_bigstep.log iptables_bigstep.nomatch) apply(simp split: ternaryvalue.split_asm add: matches_case_ternaryvalue_tuple) apply (metis in_doubt_allow_allows_Accept iptables_bigstep.nomatch matches_casesE ternaryvalue.distinct(1) ternaryvalue.distinct(5)) apply(case_tac a) apply(simp_all) apply (metis iptables_bigstep.drop iptables_bigstep.nomatch) apply (metis iptables_bigstep.log iptables_bigstep.nomatch) apply (metis iptables_bigstep.nomatch iptables_bigstep.reject) apply(simp add: good_ruleset_alt) apply(simp add: good_ruleset_alt) apply(simp add: good_ruleset_alt) apply (metis iptables_bigstep.empty iptables_bigstep.nomatch) apply(simp add: good_ruleset_alt) apply(simp add: good_ruleset_append,clarify) by (metis approximating_bigstep_to_undecided iptables_bigstep.decision iptables_bigstep.seq) lemma FinalDeny_approximating_in_doubt_allow: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ (\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalDeny \ \,\,p\ \rs, Undecided\ \ Decision FinalDeny" apply(rotate_tac 2) apply(induction rs Undecided "Decision FinalDeny" rule: approximating_bigstep_induct) apply(simp_all) apply (metis action.distinct(1) action.distinct(5) deny not_exact_match_in_doubt_allow_approx_match) apply(simp add: good_ruleset_append, clarify) apply(case_tac t) apply(simp) apply(drule(2) approximating_bigstep_undecided_to_undecided_in_doubt_allow_approx[where \=\]) apply(erule disjE) apply (metis iptables_bigstep.seq) apply (metis iptables_bigstep.decision iptables_bigstep.seq) by (metis Decision_approximating_bigstep_fun approximating_semantics_imp_fun iptables_bigstep.decision iptables_bigstep.seq) corollary FinalDenys_subseteq_in_doubt_allow: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ {p. (\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalDeny} \ {p. \,\,p\ \rs, Undecided\ \ Decision FinalDeny}" using FinalDeny_approximating_in_doubt_allow by (metis (lifting, full_types) Collect_mono) text\ If our approximating firewall (the executable version) concludes that we deny a packet, the exact semantic agrees that this packet is definitely denied! \ corollary "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ approximating_bigstep_fun (\, in_doubt_allow) p rs Undecided = (Decision FinalDeny) \ \,\,p\ \rs, Undecided\ \ Decision FinalDeny" apply(frule(1) FinalDeny_approximating_in_doubt_allow[where p=p and \=\]) apply(rule approximating_fun_imp_semantics) apply (metis good_imp_wf_ruleset) apply(simp_all) done subsection\Tactic @{const in_doubt_deny}\ lemma iptables_bigstep_undecided_to_undecided_in_doubt_deny_approx: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ \,\,p\ \rs, Undecided\ \ Undecided \ (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Undecided \ (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Decision FinalDeny" apply(rotate_tac 2) apply(induction rs Undecided Undecided rule: iptables_bigstep_induct) apply(simp_all) apply (metis approximating_bigstep.skip) apply (metis approximating_bigstep.empty approximating_bigstep.log approximating_bigstep.nomatch) apply(case_tac "a = Log") apply (metis approximating_bigstep.log approximating_bigstep.nomatch) apply(case_tac "a = Empty") apply (metis approximating_bigstep.empty approximating_bigstep.nomatch) apply(drule_tac a=a in not_exact_match_in_doubt_deny_approx_match) apply(simp_all) apply(simp add: good_ruleset_alt) apply fast apply (metis approximating_bigstep.drop approximating_bigstep.nomatch approximating_bigstep.reject) apply(frule iptables_bigstep_to_undecided) apply(simp) apply(simp add: good_ruleset_append) apply (metis (hide_lams, no_types) approximating_bigstep.decision Semantics_Ternary.seq') apply(simp add: good_ruleset_def) apply(simp add: good_ruleset_def) done lemma FinalDeny_approximating_in_doubt_deny: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ \,\,p\ \rs, Undecided\ \ Decision FinalDeny \ (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Decision FinalDeny" apply(rotate_tac 2) apply(induction rs Undecided "Decision FinalDeny" rule: iptables_bigstep_induct) apply(simp_all) apply (metis approximating_bigstep.drop approximating_bigstep.reject in_doubt_deny_denies_DropReject) apply(case_tac t) apply(simp_all) prefer 2 apply(simp add: good_ruleset_append) apply (metis approximating_bigstep.decision approximating_bigstep.seq Semantics.decisionD state.inject) - apply(thin_tac "False \ _ \ _") apply(simp add: good_ruleset_append, clarify) apply(drule(2) iptables_bigstep_undecided_to_undecided_in_doubt_deny_approx) apply(erule disjE) apply (metis approximating_bigstep.seq) apply (metis approximating_bigstep.decision Semantics_Ternary.seq') apply(simp add: good_ruleset_alt) done lemma approximating_bigstep_undecided_to_undecided_in_doubt_deny_approx: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Undecided \ \,\,p\ \rs, Undecided\ \ Undecided \ \,\,p\ \rs, Undecided\ \ Decision FinalAllow" apply(rotate_tac 2) apply(induction rs Undecided Undecided rule: approximating_bigstep_induct) apply(simp_all) apply (metis iptables_bigstep.skip) apply (metis iptables_bigstep.empty iptables_bigstep.log iptables_bigstep.nomatch) apply(simp split: ternaryvalue.split_asm add: matches_case_ternaryvalue_tuple) apply (metis in_doubt_allow_allows_Accept iptables_bigstep.nomatch matches_casesE ternaryvalue.distinct(1) ternaryvalue.distinct(5)) apply(case_tac a) apply(simp_all) apply (metis iptables_bigstep.accept iptables_bigstep.nomatch) apply (metis iptables_bigstep.log iptables_bigstep.nomatch) apply(simp add: good_ruleset_alt) apply(simp add: good_ruleset_alt) apply(simp add: good_ruleset_alt) apply (metis iptables_bigstep.empty iptables_bigstep.nomatch) apply(simp add: good_ruleset_alt) apply(simp add: good_ruleset_append,clarify) by (metis approximating_bigstep_to_undecided iptables_bigstep.decision iptables_bigstep.seq) lemma FinalAllow_approximating_in_doubt_deny: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ \,\,p\ \rs, Undecided\ \ Decision FinalAllow" apply(rotate_tac 2) apply(induction rs Undecided "Decision FinalAllow" rule: approximating_bigstep_induct) apply(simp_all) apply (metis action.distinct(1) action.distinct(5) iptables_bigstep.accept not_exact_match_in_doubt_deny_approx_match) apply(simp add: good_ruleset_append, clarify) apply(case_tac t) apply(simp) apply(drule(2) approximating_bigstep_undecided_to_undecided_in_doubt_deny_approx[where \=\]) apply(erule disjE) apply (metis iptables_bigstep.seq) apply (metis iptables_bigstep.decision iptables_bigstep.seq) by (metis Decision_approximating_bigstep_fun approximating_semantics_imp_fun iptables_bigstep.decision iptables_bigstep.seq) corollary FinalAllows_subseteq_in_doubt_deny: "matcher_agree_on_exact_matches \ \ \ good_ruleset rs \ {p. (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow} \ {p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow}" using FinalAllow_approximating_in_doubt_deny by (metis (lifting, full_types) Collect_mono) corollary new_packets_to_simple_firewall_underapproximation: defines "preprocess rs \ lower_closure (optimize_matches abstract_for_simple_firewall (lower_closure (packet_assume_new rs)))" and "newpkt p \ match_tcp_flags ipt_tcp_syn (p_tcp_flags p) \ p_tag_ctstate p = CT_New" fixes p :: "('i::len, 'pkt_ext) tagged_packet_scheme" assumes "matcher_agree_on_exact_matches \ common_matcher" and "simple_ruleset rs" shows "{p. simple_fw (to_simple_firewall (preprocess rs)) p = Decision FinalAllow \ newpkt p} \ {p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow \ newpkt p}" proof - from assms(3) have "{p. (common_matcher, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow \ newpkt p} \ {p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow \ newpkt p}" apply(drule_tac rs=rs and \=\ in FinalAllows_subseteq_in_doubt_deny) using simple_imp_good_ruleset assms(4) apply blast by blast thus ?thesis unfolding newpkt_def preprocess_def using transform_simple_fw_lower(2)[OF assms(4)] by blast qed subsection\Approximating Closures\ theorem FinalAllowClosure: assumes "matcher_agree_on_exact_matches \ \" and "good_ruleset rs" shows "{p. (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow} \ {p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow}" and "{p. \,\,p\ \rs, Undecided\ \ Decision FinalAllow} \ {p. (\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalAllow}" apply (metis FinalAllows_subseteq_in_doubt_deny assms) by (metis FinalAllows_subseteq_in_doubt_allow assms) theorem FinalDenyClosure: assumes "matcher_agree_on_exact_matches \ \" and "good_ruleset rs" shows "{p. (\, in_doubt_allow),p\ \rs, Undecided\ \\<^sub>\ Decision FinalDeny} \ {p. \,\,p\ \rs, Undecided\ \ Decision FinalDeny}" and "{p. \,\,p\ \rs, Undecided\ \ Decision FinalDeny} \ {p. (\, in_doubt_deny),p\ \rs, Undecided\ \\<^sub>\ Decision FinalDeny}" apply (metis FinalDenys_subseteq_in_doubt_allow assms) by (metis FinalDeny_approximating_in_doubt_deny assms mem_Collect_eq subsetI) subsection\Exact Embedding\ lemma LukassLemma: assumes agree: "matcher_agree_on_exact_matches \ \" and noUnknown: "(\ r \ set rs. ternary_ternary_eval (map_match_tac \ p (get_match r)) \ TernaryUnknown)" and good: "good_ruleset rs" shows "(\,\),p\ \rs, s\ \\<^sub>\ t \ \,\,p\ \rs, s\ \ t" proof - { fix t \ \if we show it for arbitrary @{term t}, we can reuse this fact for the other direction.\ assume a: "(\,\),p\ \rs, s\ \\<^sub>\ t" from a good agree noUnknown have "\,\,p\ \rs, s\ \ t" proof(induction rs s t rule: approximating_bigstep_induct) qed(auto intro: approximating_bigstep.intros iptables_bigstep.intros dest: iptables_bigstepD dest: matches_comply_exact simp: good_ruleset_append) } note 1=this { assume a: "\,\,p\ \rs, s\ \ t" obtain x where "approximating_bigstep_fun (\,\) p rs s = x" by simp with approximating_fun_imp_semantics[OF good_imp_wf_ruleset[OF good]] have x: "(\,\),p\ \rs, s\ \\<^sub>\ x" by fast with 1 have "\,\,p\ \rs, s\ \ x" by simp with a iptables_bigstep_deterministic have "x = t" by metis hence "(\,\),p\ \rs, s\ \\<^sub>\ t" using x by blast } note 2=this from 1 2 show ?thesis by blast qed text\ For rulesets without @{term Call}s, the approximating ternary semantics can perfectly simulate the Boolean semantics. \ theorem \\<^sub>m\<^sub>a\<^sub>g\<^sub>i\<^sub>c_approximating_bigstep_iff_iptables_bigstep: assumes "\r \ set rs. \c. get_action r \ Call c" shows "((\\<^sub>m\<^sub>a\<^sub>g\<^sub>i\<^sub>c \),\),p\ \rs, s\ \\<^sub>\ t \ \,\,p\ \rs, s\ \ t" apply(rule iffI) apply(induction rs s t rule: approximating_bigstep_induct) apply(auto intro: iptables_bigstep.intros simp: \\<^sub>m\<^sub>a\<^sub>g\<^sub>i\<^sub>c_matching)[7] apply(insert assms) apply(induction rs s t rule: iptables_bigstep_induct) apply(auto intro: approximating_bigstep.intros simp: \\<^sub>m\<^sub>a\<^sub>g\<^sub>i\<^sub>c_matching) done corollary \\<^sub>m\<^sub>a\<^sub>g\<^sub>i\<^sub>c_approximating_bigstep_fun_iff_iptables_bigstep: assumes "good_ruleset rs" shows "approximating_bigstep_fun (\\<^sub>m\<^sub>a\<^sub>g\<^sub>i\<^sub>c \,\) p rs s = t \ \,\,p\ \rs, s\ \ t" apply(subst approximating_semantics_iff_fun_good_ruleset[symmetric]) using assms apply simp apply(subst \\<^sub>m\<^sub>a\<^sub>g\<^sub>i\<^sub>c_approximating_bigstep_iff_iptables_bigstep[where \=\]) using assms apply (simp add: good_ruleset_def) by simp text\The function @{const optimize_primitive_univ} was only applied to the ternary semantics. It is, in fact, also correct for the Boolean semantics, assuming the @{const common_matcher}.\ lemma Semantics_optimize_primitive_univ_common_matcher: assumes "matcher_agree_on_exact_matches \ common_matcher" shows "Semantics.matches \ (optimize_primitive_univ m) p = Semantics.matches \ m p" proof - have "(max_word::16 word) = 65535" by(simp add: max_word_def) hence port_range: "\s e port. s = 0 \ e = 0xFFFF \ (port::16 word) \ 0xFFFF" by simp from assms show ?thesis apply(induction m rule: optimize_primitive_univ.induct) apply(auto elim!: matcher_agree_on_exact_matches_gammaE simp add: port_range match_ifaceAny ipset_from_cidr_0 ctstate_is_UNIV) done qed end diff --git a/thys/Neumann_Morgenstern_Utility/Neumann_Morgenstern_Utility_Theorem.thy b/thys/Neumann_Morgenstern_Utility/Neumann_Morgenstern_Utility_Theorem.thy --- a/thys/Neumann_Morgenstern_Utility/Neumann_Morgenstern_Utility_Theorem.thy +++ b/thys/Neumann_Morgenstern_Utility/Neumann_Morgenstern_Utility_Theorem.thy @@ -1,1508 +1,1508 @@ (* License: LGPL *) (* Author: Julian Parsert *) theory Neumann_Morgenstern_Utility_Theorem imports "HOL-Probability.Probability" "First_Welfare_Theorem.Utility_Functions" Lotteries begin section \ Properties of Preferences \ subsection \ Independent Preferences\ text \ Independence is sometimes called substitution \ text \ Notice how r is "added" to the right of mix-pmf and the element to the left q/p changes \ definition independent_vnm where "independent_vnm C P = (\p \ C. \q \ C. \r \ C. \(\::real) \ {0<..1}. p \[P] q \ mix_pmf \ p r \[P] mix_pmf \ q r)" lemma independent_vnmI1: assumes "(\p \ C. \q \ C. \r \ C. \\ \ {0<..1}. p \[P] q \ mix_pmf \ p r \[P] mix_pmf \ q r)" shows "independent_vnm C P" using assms independent_vnm_def by blast lemma independent_vnmI2: assumes "\p q r \. p \ C \ q \ C \ r \ C \ \ \ {0<..1} \ p \[P] q \ mix_pmf \ p r \[P] mix_pmf \ q r" shows "independent_vnm C P" by (rule independent_vnmI1, standard, standard, standard, standard, simp add: assms) (meson assms greaterThanAtMost_iff) lemma independent_vnm_alt_def: shows "independent_vnm C P \ (\p \ C. \q \ C. \r \ C. \\ \ {0<..<1}. p \[P] q \ mix_pmf \ p r \[P] mix_pmf \ q r)" (is "?L \ ?R") proof (rule iffI) assume a: "?R" have "independent_vnm C P" by(rule independent_vnmI2, simp add: a) (metis a greaterThanLessThan_iff linorder_neqE_linordered_idom not_le pmf_mix_1) then show "?L" by auto qed (simp add: independent_vnm_def) lemma independece_dest_alt: assumes "independent_vnm C P" shows "(\p \ C. \q \ C. \r \ C. \(\::real) \ {0<..1}. p \[P] q \ mix_pmf \ p r \[P] mix_pmf \ q r)" proof (standard, standard, standard, standard) fix p q r \ assume as1: "p \ C" assume as2: "q \ C" assume as3: "r \ C" assume as4: "(\::real) \ {0<..1}" then show "p \[P] q = mix_pmf \ p r \[P] mix_pmf \ q r" using as1 as2 as3 assms(1) independent_vnm_def by blast qed lemma independent_vnmD1: assumes "independent_vnm C P" shows "(\p \ C. \q \ C. \r \ C. \\ \ {0<..1}. p \[P] q \ mix_pmf \ p r \[P] mix_pmf \ q r)" using assms independent_vnm_def by blast lemma independent_vnmD2: fixes p q r \ assumes "\ \ {0<..1}" and "p \ C" and "q \ C" and "r \ C" assumes "independent_vnm C P" assumes "p \[P] q" shows "mix_pmf \ p r \[P] mix_pmf \ q r" using assms independece_dest_alt by blast lemma independent_vnmD3: fixes p q r \ assumes "\ \ {0<..1}" and "p \ C" and "q \ C" and "r \ C" assumes "independent_vnm C P" assumes "mix_pmf \ p r \[P] mix_pmf \ q r" shows "p \[P] q" using assms independece_dest_alt by blast lemma independent_vnmD4: assumes "independent_vnm C P" assumes "refl_on C P" assumes "p \ C" and "q \ C" and "r \ C" and "\ \ {0..1}" and "p \[P] q" shows "mix_pmf \ p r \[P] mix_pmf \ q r" using assms by (cases "\ = 0 \ \ \ {0<..1}",metis assms(1,2,3,4) independece_dest_alt pmf_mix_0 refl_onD, auto) lemma approx_indep_ge: assumes "x \[\] y" assumes "\ \ {0..(1::real)}" assumes rpr: "rational_preference (lotteries_on outcomes) \" and ind: "independent_vnm (lotteries_on outcomes) \" shows "\r \ lotteries_on outcomes. (mix_pmf \ y r) \[\] (mix_pmf \ x r)" proof fix r assume a: "r \ lotteries_on outcomes" (is "r \ ?lo") have clct: "y \[\] x \ independent_vnm ?lo \ \ y \ ?lo \ x \ ?lo \ r \ ?lo" by (meson a assms(1) assms(2) atLeastAtMost_iff greaterThanAtMost_iff ind preference_def rational_preference_def rpr) then have in_lo: "mix_pmf \ y r \ ?lo" "(mix_pmf \ x r) \ ?lo" by (metis assms(2) atLeastAtMost_iff greaterThanLessThan_iff less_eq_real_def mix_pmf_in_lotteries pmf_mix_0 pmf_mix_1 a)+ have "0 = \ \ 0 < \" using assms by auto then show "mix_pmf \ y r \[\] mix_pmf \ x r" using in_lo(2) rational_preference.compl rpr by (auto,blast) (meson assms(2) atLeastAtMost_iff clct greaterThanAtMost_iff independent_vnmD2) qed lemma approx_imp_approx_ind: assumes "x \[\] y" assumes "\ \ {0..(1::real)}" assumes rpr: "rational_preference (lotteries_on outcomes) \" and ind: "independent_vnm (lotteries_on outcomes) \" shows "\r \ lotteries_on outcomes. (mix_pmf \ y r) \[\] (mix_pmf \ x r)" using approx_indep_ge assms(1) assms(2) ind rpr by blast lemma geq_imp_mix_geq_right: assumes "x \[\] y" assumes rpr: "rational_preference (lotteries_on outcomes) \" assumes ind: "independent_vnm (lotteries_on outcomes) \" assumes "\ \ {0..(1::real)}" shows "(mix_pmf \ x y) \[\] y" proof - have xy_p: "x \ (lotteries_on outcomes)" "y \ (lotteries_on outcomes)" by (meson assms(1) preference.not_outside rational_preference_def rpr) (meson assms(1) preference_def rational_preference_def rpr) have "(mix_pmf \ x y) \ (lotteries_on outcomes)" (is "?mpf \ ?lot") using mix_pmf_in_lotteries [of x outcomes y \] xy_p assms(2) by (meson approx_indep_ge assms(4) ind preference.not_outside rational_preference.compl rational_preference_def) have all: "\r \ ?lot. (mix_pmf \ x r) \[\] (mix_pmf \ y r)" by (metis assms assms(2) atLeastAtMost_iff greaterThanAtMost_iff independece_dest_alt less_eq_real_def pmf_mix_0 rational_preference.compl rpr ind xy_p) thus ?thesis by (metis all assms(4) set_pmf_mix_eq xy_p(2)) qed lemma geq_imp_mix_geq_left: assumes "x \[\] y" assumes rpr: "rational_preference (lotteries_on outcomes) \" assumes ind: "independent_vnm (lotteries_on outcomes) \" assumes "\ \ {0..(1::real)}" shows "(mix_pmf \ y x) \[\] y" proof - define \ where b: "\ = 1 - \" have "\ \ {0..1}" using assms(4) b by auto then have "mix_pmf \ x y \[\] y" using geq_imp_mix_geq_right[OF assms] assms(1) geq_imp_mix_geq_right ind rpr by blast moreover have "mix_pmf \ x y = mix_pmf \ y x" by (metis assms(4) b pmf_inverse_switch_eqals) ultimately show ?thesis by simp qed lemma sg_imp_mix_sg: assumes "x \[\] y" assumes rpr: "rational_preference (lotteries_on outcomes) \" assumes ind: "independent_vnm (lotteries_on outcomes) \" assumes "\ \ {0<..(1::real)}" shows "(mix_pmf \ x y) \[\] y" proof - have xy_p: "x \ (lotteries_on outcomes)" "y \ (lotteries_on outcomes)" by (meson assms(1) preference.not_outside rational_preference_def rpr) (meson assms(1) preference_def rational_preference_def rpr) have "(mix_pmf \ x y) \ (lotteries_on outcomes)" (is "?mpf \ ?lot") using mix_pmf_in_lotteries [of x outcomes y \] xy_p assms(2) using assms(4) by fastforce have all: "\r \ ?lot. (mix_pmf \ x r) \[\] (mix_pmf \ y r)" by (metis assms(1,3,4) independece_dest_alt ind xy_p) have "(mix_pmf \ x y) \[\] y" by (metis all assms(4) atLeastAtMost_iff greaterThanAtMost_iff less_eq_real_def set_pmf_mix_eq xy_p(2)) have all2: "\r \ ?lot. (mix_pmf \ x r) \[\] (mix_pmf \ y r)" using assms(1) assms(4) ind independece_dest_alt xy_p(1) xy_p(2) by blast then show ?thesis by (metis assms(4) atLeastAtMost_iff greaterThanAtMost_iff less_eq_real_def set_pmf_mix_eq xy_p(2)) qed subsection \ Continuity \ text \ Continuity is sometimes called Archimedean Axiom\ definition continuous_vnm where "continuous_vnm C P = (\p \ C. \q \ C. \r \ C. p \[P] q \ q \[P] r \ (\\ \ {0..1}. (mix_pmf \ p r) \[P] q))" lemma continuous_vnmD: assumes "continuous_vnm C P" shows "(\p \ C. \q \ C. \r \ C. p \[P] q \ q \[P] r \ (\\ \ {0..1}. (mix_pmf \ p r) \[P] q))" using continuous_vnm_def assms by blast lemma continuous_vnmI: assumes "\p q r. p \ C \ q \ C \ r \ C \ p \[P] q \ q \[P] r \ \\ \ {0..1}. (mix_pmf \ p r) \[P] q" shows "continuous_vnm C P" by (simp add: assms continuous_vnm_def) lemma mix_in_lot: assumes "x \ lotteries_on outcomes" and "y \ lotteries_on outcomes" and "\ \ {0..1}" shows "(mix_pmf \ x y) \ lotteries_on outcomes" using assms(1) assms(2) assms(3) less_eq_real_def mix_pmf_in_lotteries by fastforce lemma non_unique_continuous_unfolding: assumes cnt: "continuous_vnm (lotteries_on outcomes) \" assumes "rational_preference (lotteries_on outcomes) \" assumes "p \[\] q" and "q \[\] r" and "p \[\] r" shows "\\ \ {0..1}. q \[\] mix_pmf \ p r" using assms(1) assms(2) cnt continuous_vnmD assms proof - have "\p q. p\ (lotteries_on outcomes) \ q \ (lotteries_on outcomes) \ p \[\] q \ q \[\] p" using assms rational_preference.compl[of "lotteries_on outcomes" \] by (metis (no_types, hide_lams) preference_def rational_preference_def) then show ?thesis using continuous_vnmD[OF assms(1)] by (metis assms(3) assms(4)) qed section \ System U start, as per vNM\ text \ These are the first two assumptions which we use to derive the first results. We assume rationality and independence. In this system U the von-Neumann-Morgenstern Utility Theorem is proven. \ context fixes outcomes :: "'a set" fixes \ assumes rpr: "rational_preference (lotteries_on outcomes) \" assumes ind: "independent_vnm (lotteries_on outcomes) \" begin abbreviation "\

\ lotteries_on outcomes" lemma relation_in_carrier: "x \[\] y \ x \ \

\ y \ \

" by (meson preference_def rational_preference_def rpr) lemma mix_pmf_preferred_independence: assumes "r \ \

" and "\ \ {0..1}" assumes "p \[\] q" shows "mix_pmf \ p r \[\] mix_pmf \ q r" using ind by (metis relation_in_carrier antisym_conv1 assms atLeastAtMost_iff greaterThanAtMost_iff independece_dest_alt pmf_mix_0 rational_preference.no_better_thansubset_rel rpr subsetI) lemma mix_pmf_strict_preferred_independence: assumes "r \ \

" and "\ \ {0<..1}" assumes "p \[\] q" shows "mix_pmf \ p r \[\] mix_pmf \ q r" by (meson assms(1) assms(2) assms(3) ind independent_vnmD2 independent_vnmD3 relation_in_carrier) lemma mix_pmf_preferred_independence_rev: assumes "p \ \

" and "q \ \

" and "r \ \

" and "\ \ {0<..1}" assumes "mix_pmf \ p r \[\] mix_pmf \ q r" shows "p \[\] q" proof - have "mix_pmf \ p r \ \

" using assms mix_in_lot relation_in_carrier by blast moreover have "mix_pmf \ q r \ \

" using assms mix_in_lot assms(2) relation_in_carrier by blast ultimately show ?thesis using ind independent_vnmD3[of \ p \

q r \] assms by blast qed lemma x_sg_y_sg_mpmf_right: assumes "x \[\] y" assumes "b \ {0<..(1::real)}" shows "x \[\] mix_pmf b y x" proof - consider "b = 1" | "b \ 1" by blast then show ?thesis proof (cases) case 2 have sg: "(mix_pmf b x y) \[\] y" using assms(1) assms(2) assms ind rpr sg_imp_mix_sg "2" by fastforce have "mix_pmf b x y \ \

" by (meson sg preference_def rational_preference_def rpr) have "mix_pmf b x x \ \

" using relation_in_carrier assms(2) mix_in_lot assms by fastforce have "b \ {0<..<1}" using "2" assms(2) by auto have "mix_pmf b x x \[\] mix_pmf b y x" using mix_pmf_preferred_independence[of x b] assms by (meson \b \ {0<..<1}\ greaterThanAtMost_iff greaterThanLessThan_iff ind independece_dest_alt less_eq_real_def preference_def rational_preference.axioms(1) relation_in_carrier rpr) then show ?thesis using mix_pmf_preferred_independence by (metis assms(2) atLeastAtMost_iff greaterThanAtMost_iff less_eq_real_def set_pmf_mix_eq) qed (simp add: assms(1)) qed lemma neumann_3B_b: assumes "u \[\] v" assumes "\ \ {0<..<1}" shows "u \[\] mix_pmf \ u v" proof - have *: "preorder_on \

\ \ rational_preference_axioms \

\" by (metis (no_types) preference_def rational_preference_def rpr) have "1 - \ \ {0<..1}" using assms(2) by auto then show ?thesis using * assms by (metis atLeastAtMost_iff greaterThanLessThan_iff less_eq_real_def pmf_inverse_switch_eqals x_sg_y_sg_mpmf_right) qed lemma neumann_3B_b_non_strict: assumes "u \[\] v" assumes "\ \ {0..1}" shows "u \[\] mix_pmf \ u v" proof - have f2: "mix_pmf \ (u::'a pmf) v = mix_pmf (1 - \) v u" using pmf_inverse_switch_eqals assms(2) by auto have "1 - \ \ {0..1}" using assms(2) by force then show ?thesis using f2 relation_in_carrier by (metis (no_types) assms(1) mix_pmf_preferred_independence set_pmf_mix_eq) qed lemma greater_mix_pmf_greater_step_1_aux: assumes "v \[\] u" assumes "\ \ {0<..<(1::real)}" and "\ \ {0<..<(1::real)}" assumes "\ > \" shows "(mix_pmf \ v u) \[\] (mix_pmf \ v u)" proof - define t where t: "t = mix_pmf \ v u" obtain \ where g: "\ = \ * \" by (metis assms(2) assms(4) greaterThanLessThan_iff mult.commute nonzero_eq_divide_eq not_less_iff_gr_or_eq) have g1: "\ > 0 \ \ < 1" by (metis (full_types) assms(2) assms(4) g greaterThanLessThan_iff less_trans mult.right_neutral mult_less_cancel_left_pos not_le sgn_le_0_iff sgn_pos zero_le_one zero_le_sgn_iff zero_less_mult_iff) have t_in: "mix_pmf \ v u \ \

" by (meson assms(1) assms(3) mix_pmf_in_lotteries preference_def rational_preference_def rpr) have "v \[\] mix_pmf (1 - \) v u" using x_sg_y_sg_mpmf_right[of u v "1-\"] assms by (metis atLeastAtMost_iff greaterThanAtMost_iff greaterThanLessThan_iff less_eq_real_def pmf_inverse_switch_eqals x_sg_y_sg_mpmf_right) have "t \[\] u" using assms(1) assms(3) ind rpr sg_imp_mix_sg t by fastforce hence t_s: "t \[\] (mix_pmf \ t u)" proof - have "(mix_pmf \ t u) \ \

" by (metis assms(1) assms(3) atLeastAtMost_iff g1 mix_in_lot mix_pmf_in_lotteries not_less order.asym preference_def rational_preference_def rpr t) have "t \[\] mix_pmf \ (mix_pmf \ v u) u" using neumann_3B_b[of t u \] assms t g1 by (meson greaterThanAtMost_iff greaterThanLessThan_iff ind less_eq_real_def rpr sg_imp_mix_sg) thus ?thesis using t by blast qed from product_mix_pmf_prob_distrib[of _ \ v u] assms have "mix_pmf \ v u \[\] mix_pmf \ v u" by (metis t_s atLeastAtMost_iff g g1 greaterThanLessThan_iff less_eq_real_def mult.commute t) then show ?thesis by blast qed section \ This lemma is in called step 1 in literature. In Von Neumann and Morgenstern's book this is A:A (albeit more general) \ lemma step_1_most_general: assumes "x \[\] y" assumes "\ \ {0..(1::real)}" and "\ \ {0..(1::real)}" assumes "\ > \" shows "(mix_pmf \ x y) \[\] (mix_pmf \ x y)" proof - consider (ex) "\ = 1 \ \ = 0" | (m) "\ \ 1 \ \ \ 0" by blast then show ?thesis proof (cases) case m consider "\ = 0" | "\ \ 0" by blast then show ?thesis proof (cases) case 1 then show ?thesis using assms(1) assms(2) assms(4) ind rpr sg_imp_mix_sg by fastforce next case 2 let ?d = "(\/\)" have sg: "(mix_pmf \ x y) \[\] y" using assms(1) assms(2) assms(3) assms(4) ind rpr sg_imp_mix_sg by fastforce have a: "\ > 0" using assms(3) assms(4) by auto then have div_in: "?d \ {0<..1}" using assms(3) assms(4) 2 by auto have mx_p: "(mix_pmf \ x y) \ \

" by (meson sg preference_def rational_preference_def rpr) have y_P: "y \ \

" by (meson assms(1) preference_def rational_preference_def rpr) hence "(mix_pmf ?d (mix_pmf \ x y) y) \ \

" using div_in mx_p by (simp add: mix_in_lot) have " mix_pmf \ (mix_pmf \ x y) y \[\] y" using sg_imp_mix_sg[of "(mix_pmf \ x y)" y \ outcomes \] sg div_in rpr ind a assms(2) "2" assms(3) by auto have al1: "\r \ \

. (mix_pmf \ x r) \[\] (mix_pmf \ y r)" by (meson a assms(1) assms(2) atLeastAtMost_iff greaterThanAtMost_iff ind independece_dest_alt preference.not_outside rational_preference_def rpr y_P) then show ?thesis using greater_mix_pmf_greater_step_1_aux assms by (metis a div_in divide_less_eq_1_pos greaterThanAtMost_iff greaterThanLessThan_iff mix_pmf_comp_with_dif_equiv neumann_3B_b sg) qed qed (simp add: assms(1)) qed text \ Kreps refers to this lemma as 5.6 c. The lemma after that is also significant.\ lemma approx_remains_after_same_comp: assumes "p \[\] q" and "r \ \

" and "\ \ {0..1}" shows "mix_pmf \ p r \[\] mix_pmf \ q r" using approx_indep_ge assms(1) assms(2) assms(3) ind rpr by blast text \ This lemma is the symmetric version of the previous lemma. This lemma is never mentioned in literature anywhere. Even though it looks trivial now, due to the asymmetric nature of the independence axiom, it is not so trivial, and definitely worth mentioning. \ lemma approx_remains_after_same_comp_left: assumes "p \[\] q" and "r \ \

" and "\ \ {0..1}" shows "mix_pmf \ r p \[\] mix_pmf \ r q" proof - have 1: "\ \ 1 \ \ \ 0" "1 - \ \ {0..1}" using assms(3) by auto+ have fst: "mix_pmf \ r p \[\] mix_pmf (1-\) p r" using assms by (metis mix_in_lot pmf_inverse_switch_eqals rational_preference.compl relation_in_carrier rpr) moreover have "mix_pmf \ r p \[\] mix_pmf \ r q" using approx_remains_after_same_comp[of _ _ _ \] pmf_inverse_switch_eqals[of \ p q] 1 pmf_inverse_switch_eqals rpr mix_pmf_preferred_independence[of _ \ _ _] by (metis assms(1) assms(2) assms(3) mix_pmf_preferred_independence) thus ?thesis by blast qed lemma mix_of_preferred_is_preferred: assumes "p \[\] w" assumes "q \[\] w" assumes "\ \ {0..1}" shows "mix_pmf \ p q \[\] w" proof - consider "p \[\] q" | "q \[\] p" using rpr assms(1) assms(2) rational_preference.compl relation_in_carrier by blast then show ?thesis proof (cases) case 1 have "mix_pmf \ p q \[\] q" using "1" assms(3) geq_imp_mix_geq_right ind rpr by blast moreover have "q \[\] w" using assms by auto ultimately show ?thesis using rpr preference.transitivity[of \

\] by (meson rational_preference_def transE) next case 2 have "mix_pmf \ p q \[\] p" using "2" assms geq_imp_mix_geq_left ind rpr by blast moreover have "p \[\] w" using assms by auto ultimately show ?thesis using rpr preference.transitivity[of \

\] by (meson rational_preference_def transE) qed qed lemma mix_of_not_preferred_is_not_preferred: assumes "w \[\] p" assumes "w \[\] q" assumes "\ \ {0..1}" shows "w \[\] mix_pmf \ p q" proof - consider "p \[\] q" | "q \[\] p" using rpr assms(1) assms(2) rational_preference.compl relation_in_carrier by blast then show ?thesis proof (cases) case 1 moreover have "p \[\] mix_pmf \ p q" using assms(3) neumann_3B_b_non_strict calculation by blast moreover show ?thesis using rpr preference.transitivity[of \

\] by (meson assms(1) calculation(2) rational_preference_def transE) next case 2 moreover have "q \[\] mix_pmf \ p q" using assms(3) neumann_3B_b_non_strict calculation by (metis mix_pmf_preferred_independence relation_in_carrier set_pmf_mix_eq) moreover show ?thesis using rpr preference.transitivity[of \

\] by (meson assms(2) calculation(2) rational_preference_def transE) qed qed private definition degenerate_lotteries where "degenerate_lotteries = {x \ \

. card (set_pmf x) = 1}" private definition best where "best = {x \ \

. (\y \ \

. x \[\] y)}" private definition worst where "worst = {x \ \

. (\y \ \

. y \[\] x)}" lemma degenerate_total: "\e \ degenerate_lotteries. \m \ \

. e \[\] m \ m \[\] e" using degenerate_lotteries_def rational_preference.compl rpr by fastforce lemma degen_outcome_cardinalities: "card degenerate_lotteries = card outcomes" using card_degen_lotteries_equals_outcomes degenerate_lotteries_def by auto lemma degenerate_lots_subset_all: "degenerate_lotteries \ \

" by (simp add: degenerate_lotteries_def) lemma alt_definition_of_degenerate_lotteries[iff]: "{return_pmf x |x. x\ outcomes} = degenerate_lotteries" proof (standard, goal_cases) case 1 have "\x \ {return_pmf x |x. x \ outcomes}. x \ degenerate_lotteries" proof fix x assume a: "x \ {return_pmf x |x. x \ outcomes}" then have "card (set_pmf x) = 1" by auto moreover have "set_pmf x \ outcomes" using a set_pmf_subset_singleton by auto moreover have "x \ \

" by (simp add: lotteries_on_def calculation) ultimately show "x \ degenerate_lotteries" by (simp add: degenerate_lotteries_def) qed then show ?case by blast next case 2 have "\x \ degenerate_lotteries. x \ {return_pmf x |x. x \ outcomes}" proof fix x assume a: "x \ degenerate_lotteries" hence "card (set_pmf x) = 1" using degenerate_lotteries_def by blast moreover have "set_pmf x \ outcomes" by (meson a degenerate_lots_subset_all subset_iff support_in_outcomes) moreover obtain e where "{e} = set_pmf x" using calculation by (metis card_1_singletonE) moreover have "e \ outcomes" using calculation(2) calculation(3) by blast moreover have "x = return_pmf e" - using calculation(3) set_pmf_subset_singleton by fastforce + using calculation(3) set_pmf_subset_singleton by fast ultimately show "x \ {return_pmf x |x. x \ outcomes}" by blast qed then show ?case by blast qed lemma best_indifferent: "\x \ best. \y \ best. x \[\] y" by (simp add: best_def) lemma worst_indifferent: "\x \ worst. \y \ worst. x \[\] y" by (simp add: worst_def) lemma best_worst_indiff_all_indiff: assumes "b \ best" and "w \ worst" and "b \[\] w" shows "\e \ \

. e \[\] w" "\e \ \

. e \[\] b" proof - show "\e \ \

. e \[\] w" proof (standard) fix e assume a: "e \ \

" then have "b \[\] e" using a best_def assms by blast moreover have "e \[\] w" using a assms worst_def by auto moreover have "b \[\] e" by (simp add: calculation(1)) moreover show "e \[\] w" proof (rule ccontr) assume "\ e \[\] w" then consider "e \[\] w" | "w \[\] e" by (simp add: calculation(2)) then show False proof (cases) case 2 then show ?thesis using calculation(2) by blast qed (meson assms(3) calculation(1) rational_preference.strict_is_neg_transitive relation_in_carrier rpr) qed qed then show "\e\local.\

. e \[\] b" using assms by (meson rational_preference.compl rational_preference.strict_is_neg_transitive relation_in_carrier rpr) qed text \ Like Step 1 most general but with IFF. \ lemma mix_pmf_pref_iff_more_likely [iff]: assumes "b \[\] w" assumes "\ \ {0..1}" and "\ \ {0..1}" shows "\ > \ \ mix_pmf \ b w \[\] mix_pmf \ b w" (is "?L \ ?R") using assms step_1_most_general[of b w \ \] by (metis linorder_neqE_linordered_idom step_1_most_general) lemma better_worse_good_mix_preferred[iff]: assumes "b \[\] w" assumes "\ \ {0..1}" and "\ \ {0..1}" assumes "\ \ \" shows "mix_pmf \ b w \[\] mix_pmf \ b w" proof- have "(0::real) \ 1" by simp then show ?thesis by (metis (no_types) assms assms(1) assms(2) assms(3) atLeastAtMost_iff less_eq_real_def mix_of_not_preferred_is_not_preferred mix_of_preferred_is_preferred mix_pmf_preferred_independence pmf_mix_0 relation_in_carrier step_1_most_general) qed subsection \ Add finiteness and non emptyness of outcomes \ context assumes fnt: "finite outcomes" assumes nempty: "outcomes \ {}" begin lemma finite_degenerate_lotteries: "finite degenerate_lotteries" using degen_outcome_cardinalities fnt nempty by fastforce lemma degenerate_has_max_preferred: "{x \ degenerate_lotteries. (\y \ degenerate_lotteries. x \[\] y)} \ {}" (is "?l \ {}") proof assume a: "?l = {}" let ?DG = "degenerate_lotteries" obtain R where R: "rational_preference ?DG R" "R \ \" using degenerate_lots_subset_all rational_preference.all_carrier_ex_sub_rel rpr by blast then have "\e \ ?DG. \e' \ ?DG. e \[\] e'" by (metis R(1) R(2) card_0_eq degen_outcome_cardinalities finite_degenerate_lotteries fnt nempty subset_eq rational_preference.finite_nonempty_carrier_has_maximum ) then show False using a by auto qed lemma degenerate_has_min_preferred: "{x \ degenerate_lotteries. (\y \ degenerate_lotteries. y \[\] x)} \ {}" (is "?l \ {}") proof assume a: "?l = {}" let ?DG = "degenerate_lotteries" obtain R where R: "rational_preference ?DG R" "R \ \" using degenerate_lots_subset_all rational_preference.all_carrier_ex_sub_rel rpr by blast have "\e \ ?DG. \e' \ ?DG. e' \[\] e" by (metis R(1) R(2) card_0_eq degen_outcome_cardinalities finite_degenerate_lotteries fnt nempty subset_eq rational_preference.finite_nonempty_carrier_has_minimum ) then show False using a by auto qed lemma exists_best_degenerate: "\x \ degenerate_lotteries. \y \ degenerate_lotteries. x \[\] y" using degenerate_has_max_preferred by blast lemma exists_worst_degenerate: "\x \ degenerate_lotteries. \y \ degenerate_lotteries. y \[\] x" using degenerate_has_min_preferred by blast lemma best_degenerate_in_best_overall: "\x \ degenerate_lotteries. \y \ \

. x \[\] y" proof - obtain b where b: "b \ degenerate_lotteries" "\y \ degenerate_lotteries. b \[\] y" using exists_best_degenerate by blast have asm: "finite outcomes" "set_pmf b \ outcomes" by (simp add: fnt) (meson b(1) degenerate_lots_subset_all subset_iff support_in_outcomes) obtain B where B: "set_pmf b = {B}" using b card_1_singletonE degenerate_lotteries_def by blast have deg: "\d\outcomes. b \[\] return_pmf d" using alt_definition_of_degenerate_lotteries b(2) by blast define P where "P = (\p. p \ \

\ return_pmf B \[\] p)" have "P p" for p proof - consider "set_pmf p \ outcomes" | "\set_pmf p \ outcomes" by blast then show ?thesis proof (cases) case 1 have "finite outcomes" "set_pmf p \ outcomes" by (auto simp: 1 asm) then show ?thesis proof (induct rule: pmf_mix_induct') case (degenerate x) then show ?case using B P_def deg set_pmf_subset_singleton by fastforce qed (simp add: P_def lotteries_on_def mix_of_not_preferred_is_not_preferred mix_of_not_preferred_is_not_preferred[of b p q a]) qed (simp add: lotteries_on_def P_def) qed moreover have "\e \ \

. b \[\] e" using calculation B P_def set_pmf_subset_singleton by fastforce ultimately show ?thesis using b degenerate_lots_subset_all by blast qed lemma worst_degenerate_in_worst_overall: "\x \ degenerate_lotteries. \y \ \

. y \[\] x" proof - obtain b where b: "b \ degenerate_lotteries" "\y \ degenerate_lotteries. y \[\] b" using exists_worst_degenerate by blast have asm: "finite outcomes" "set_pmf b \ outcomes" by (simp add: fnt) (meson b(1) degenerate_lots_subset_all subset_iff support_in_outcomes) obtain B where B: "set_pmf b = {B}" using b card_1_singletonE degenerate_lotteries_def by blast have deg: "\d\outcomes. return_pmf d \[\] b" using alt_definition_of_degenerate_lotteries b(2) by blast define P where "P = (\p. p \ \

\ p \[\] return_pmf B)" have "P p" for p proof - consider "set_pmf p \ outcomes" | "\set_pmf p \ outcomes" by blast then show ?thesis proof (cases) case 1 have "finite outcomes" "set_pmf p \ outcomes" by (auto simp: 1 asm) then show ?thesis proof (induct rule: pmf_mix_induct') case (degenerate x) then show ?case using B P_def deg set_pmf_subset_singleton by fastforce next qed (simp add: P_def lotteries_on_def mix_of_preferred_is_preferred mix_of_not_preferred_is_not_preferred[of b p]) qed (simp add: lotteries_on_def P_def) qed moreover have "\e \ \

. e \[\] b" using calculation B P_def set_pmf_subset_singleton by fastforce ultimately show ?thesis using b degenerate_lots_subset_all by blast qed lemma overall_best_nonempty: "best \ {}" using best_def best_degenerate_in_best_overall degenerate_lots_subset_all by blast lemma overall_worst_nonempty: "worst \ {}" using degenerate_lots_subset_all worst_def worst_degenerate_in_worst_overall by auto lemma trans_approx: assumes "x\[\] y" and " y \[\] z" shows "x \[\] z" using preference.indiff_trans[of \

\ x y z] assms rpr rational_preference_def by blast text \ First EXPLICIT use of the axiom of choice \ private definition some_best where "some_best = (SOME x. x \ degenerate_lotteries \ x \ best)" private definition some_worst where "some_worst = (SOME x. x \ degenerate_lotteries \ x \ worst)" private definition my_U :: "'a pmf \ real" where "my_U p = (SOME \. \\{0..1} \ p \[\] mix_pmf \ some_best some_worst)" lemma exists_best_and_degenerate: "degenerate_lotteries \ best \ {}" using best_def best_degenerate_in_best_overall degenerate_lots_subset_all by blast lemma exists_worst_and_degenerate: "degenerate_lotteries \ worst \ {}" using worst_def worst_degenerate_in_worst_overall degenerate_lots_subset_all by blast lemma some_best_in_best: "some_best \ best" using exists_best_and_degenerate some_best_def by (metis (mono_tags, lifting) Int_emptyI some_eq_ex) lemma some_worst_in_worst: "some_worst \ worst" using exists_worst_and_degenerate some_worst_def by (metis (mono_tags, lifting) Int_emptyI some_eq_ex) lemma best_always_at_least_as_good_mix: assumes "\ \ {0..1}" and "p \ \

" shows "mix_pmf \ some_best p \[\] p" using assms(1) assms(2) best_def mix_of_preferred_is_preferred rational_preference.compl rpr some_best_in_best by fastforce lemma geq_mix_imp_weak_pref: assumes "\ \ {0..1}" and "\ \ {0..1}" assumes "\ \ \" shows "mix_pmf \ some_best some_worst \[\] mix_pmf \ some_best some_worst" using assms(1) assms(2) assms(3) best_def some_best_in_best some_worst_in_worst worst_def by auto lemma gamma_inverse: assumes "\ \ {0<..<1}" and "\ \ {0<..<1}" shows "(1::real) - (\ - \) / (1 - \) = (1 - \) / (1 - \)" proof - have "1 - (\ - \) / (1 - \) = (1 - \)/(1 - \) - (\ - \) / (1 - \)" using assms(2) by auto also have "... = (1 - \ - (\ - \)) / (1 - \)" by (metis diff_divide_distrib) also have "... = (1 - \) / (1 - \)" by simp finally show ?thesis . qed lemma all_mix_pmf_indiff_indiff_best_worst: assumes "l \ \

" assumes "b \ best" assumes "w \ worst" assumes "b \[\] w" shows "\\ \{0..1}. l \[\] mix_pmf \ b w" by (meson assms best_worst_indiff_all_indiff(1) mix_of_preferred_is_preferred best_worst_indiff_all_indiff(2) mix_of_not_preferred_is_not_preferred) lemma indiff_imp_same_utility_value: assumes "some_best \[\] some_worst" assumes "\ \ {0..1}" assumes "\ \ {0..1}" assumes "mix_pmf \ some_best some_worst \[\] mix_pmf \ some_best some_worst" shows "\ = \" using assms(1) assms(2) assms(3) assms(4) linorder_neqE_linordered_idom by blast lemma leq_mix_imp_weak_inferior: assumes "some_best \[\] some_worst" assumes "\ \ {0..1}" and "\ \ {0..1}" assumes "mix_pmf \ some_best some_worst \[\] mix_pmf \ some_best some_worst" shows "\ \ \" proof - have *: "mix_pmf \ some_best some_worst \[\] mix_pmf \ some_best some_worst \ \ \ \" using assms(1) assms(2) assms(3) indiff_imp_same_utility_value by blast consider "mix_pmf \ some_best some_worst \[\] mix_pmf \ some_best some_worst" | "mix_pmf \ some_best some_worst \[\] mix_pmf \ some_best some_worst" using assms(4) by blast then show ?thesis by(cases) (meson assms(2) assms(3) geq_mix_imp_weak_pref le_cases *)+ qed lemma ge_mix_pmf_preferred: assumes "x \[\] y" assumes "\ \ {0..1}" and "\ \ {0..1}" assumes "\ \ \" shows "(mix_pmf \ x y) \[\] (mix_pmf \ x y)" using assms(1) assms(2) assms(3) assms(4) by blast subsection \ Add continuity to assumptions \ context assumes cnt: "continuous_vnm (lotteries_on outcomes) \" begin text \ In Literature this is referred to as step 2. \ lemma step_2_unique_continuous_unfolding: assumes "p \[\] q" and "q \[\] r" and "p \[\] r" shows "\!\ \ {0..1}. q \[\] mix_pmf \ p r" proof (rule ccontr) assume neg_a: "\!\. \ \ {0..1} \ q \[\] mix_pmf \ p r" have "\\ \ {0..1}. q \[\] mix_pmf \ p r" using non_unique_continuous_unfolding[of outcomes \ p q r] assms cnt rpr by blast then obtain \ \ :: real where a_b: "\\{0..1}" "\ \{0..1}" "q \[\] mix_pmf \ p r" "q \[\] mix_pmf \ p r" "\ \ \" using neg_a by blast consider "\ > \" | "\ > \" using a_b by linarith then show False proof (cases) case 1 with step_1_most_general[of p r \ \] assms have "mix_pmf \ p r \[\] mix_pmf \ p r" using a_b(1) a_b(2) by blast then show ?thesis using a_b by (meson rational_preference.strict_is_neg_transitive relation_in_carrier rpr) next case 2 with step_1_most_general[of p r \ \] assms have "mix_pmf \ p r \[\]mix_pmf \ p r" using a_b(1) a_b(2) by blast then show ?thesis using a_b by (meson rational_preference.strict_is_neg_transitive relation_in_carrier rpr) qed qed text \ These folowing two lemmas are referred to sometimes called step 2. \ lemma create_unique_indiff_using_distinct_best_worst: assumes "l \ \

" assumes "b \ best" assumes "w \ worst" assumes "b \[\] w" shows "\!\ \{0..1}. l \[\] mix_pmf \ b w" proof - have "b \[\] l" using best_def using assms by blast moreover have "l \[\] w" using worst_def assms by blast ultimately show "\!\\{0..1}. l \[\] mix_pmf \ b w" using step_2_unique_continuous_unfolding[of b l w] assms by linarith qed lemma exists_element_bw_mix_is_approx: assumes "l \ \

" assumes "b \ best" assumes "w \ worst" shows "\\ \{0..1}. l \[\] mix_pmf \ b w" proof - consider "b \[\] w" | "b \[\] w" using assms(2) assms(3) best_def worst_def by auto then show ?thesis proof (cases) case 1 then show ?thesis using create_unique_indiff_using_distinct_best_worst assms by blast qed (auto simp: all_mix_pmf_indiff_indiff_best_worst assms) qed lemma my_U_is_defined: assumes "p \ \

" shows "my_U p \ {0..1}" "p \[\] mix_pmf (my_U p) some_best some_worst" proof - have "some_best \ best" by (simp add: some_best_in_best) moreover have "some_worst \ worst" by (simp add: some_worst_in_worst) with exists_element_bw_mix_is_approx[of p "some_best" "some_worst"] calculation assms have e: "\\\{0..1}. p \[\] mix_pmf \ some_best some_worst" by blast then show "my_U p \ {0..1}" by (metis (mono_tags, lifting) my_U_def someI_ex) show "p \[\] mix_pmf (my_U p) some_best some_worst" by (metis (mono_tags, lifting) e my_U_def someI_ex) qed lemma weak_pref_mix_with_my_U_weak_pref: assumes "p \[\] q" shows "mix_pmf (my_U p) some_best some_worst \[\] mix_pmf (my_U q) some_best some_worst" by (meson assms my_U_is_defined(2) relation_in_carrier rpr rational_preference.weak_is_transitive) lemma preferred_greater_my_U: assumes "p \ \

" and "q \ \

" assumes "mix_pmf (my_U p) some_best some_worst \[\] mix_pmf (my_U q) some_best some_worst" shows "my_U p > my_U q" proof (rule ccontr) assume "\ my_U p > my_U q" then consider "my_U p = my_U q" | "my_U p < my_U q" by linarith then show False proof (cases) case 1 then have "mix_pmf (my_U p) some_best some_worst \[\] mix_pmf (my_U q) some_best some_worst" using assms by auto then show ?thesis using assms by blast next case 2 moreover have "my_U q \ {0..1}" using assms(2) my_U_is_defined(1) by blast moreover have "my_U p \ {0..1}" using assms(1) my_U_is_defined(1) by blast moreover have "mix_pmf (my_U q) some_best some_worst \[\] mix_pmf (my_U p) some_best some_worst" using calculation geq_mix_imp_weak_pref by auto then show ?thesis using assms by blast qed qed lemma geq_my_U_imp_weak_preference: assumes "p \ \

" and "q \ \

" assumes "some_best \[\] some_worst" assumes "my_U p \ my_U q" shows "p \[\] q" proof - have p_q: "my_U p \ {0..1}" "my_U q \ {0..1}" using assms my_U_is_defined(1) by blast+ with ge_mix_pmf_preferred[of "some_best" "some_worst" "my_U p" "my_U q"] p_q assms(1) assms(3) assms(4) have "mix_pmf (my_U p) some_best some_worst \[\] mix_pmf (my_U q) some_best some_worst" by blast consider "my_U p = my_U q" | "my_U p > my_U q" using assms by linarith then show ?thesis proof (cases) case 2 then show ?thesis by (meson assms(1) assms(2) assms(3) p_q(1) p_q(2) rational_preference.compl rpr step_1_most_general weak_pref_mix_with_my_U_weak_pref) qed (metis assms(1) assms(2) my_U_is_defined(2) trans_approx) qed lemma my_U_represents_pref: assumes "some_best \[\] some_worst" assumes "p \ \

" and "q \ \

" shows "p \[\] q \ my_U p \ my_U q" (is "?L \ ?R") proof - have p_def: "my_U p\ {0..1}" "my_U q \ {0..1}" using assms my_U_is_defined by blast+ show ?thesis proof assume a: ?L hence "mix_pmf (my_U p) some_best some_worst \[\] mix_pmf (my_U q) some_best some_worst" using weak_pref_mix_with_my_U_weak_pref by auto then show ?R using leq_mix_imp_weak_inferior[of "my_U p" "my_U q"] p_def a assms(1) leq_mix_imp_weak_inferior by blast next assume ?R then show ?L using geq_my_U_imp_weak_preference using assms(1) assms(2) assms(3) by blast qed qed lemma first_iff_u_greater_strict_preff: assumes "p \ \

" and "q \ \

" assumes "some_best \[\] some_worst" shows "my_U p > my_U q \ mix_pmf (my_U p) some_best some_worst \[\] mix_pmf (my_U q) some_best some_worst" proof assume a: "my_U p > my_U q" have "my_U p \ {0..1}" "my_U q \ {0..1}" using assms my_U_is_defined(1) by blast+ then show "mix_pmf (my_U p) some_best some_worst \[\] mix_pmf (my_U q) some_best some_worst" using a assms(3) by blast next assume a: "mix_pmf (my_U p) some_best some_worst \[\] mix_pmf (my_U q) some_best some_worst" have "my_U p \ {0..1}" "my_U q \ {0..1}" using assms my_U_is_defined(1) by blast+ then show "my_U p > my_U q " using preferred_greater_my_U[of p q] assms a by blast qed lemma second_iff_calib_mix_pref_strict_pref: assumes "p \ \

" and "q \ \

" assumes "some_best \[\] some_worst" shows "mix_pmf (my_U p) some_best some_worst \[\] mix_pmf (my_U q) some_best some_worst \ p \[\] q" proof assume a: "mix_pmf (my_U p) some_best some_worst \[\] mix_pmf (my_U q) some_best some_worst" have "my_U p \ {0..1}" "my_U q \ {0..1}" using assms my_U_is_defined(1) by blast+ then show "p \[\] q" using a assms(3) assms(1) assms(2) geq_my_U_imp_weak_preference leq_mix_imp_weak_inferior weak_pref_mix_with_my_U_weak_pref by blast next assume a: "p \[\] q" have "my_U p \ {0..1}" "my_U q \ {0..1}" using assms my_U_is_defined(1) by blast+ then show "mix_pmf (my_U p) some_best some_worst \[\] mix_pmf (my_U q) some_best some_worst" using a assms(1) assms(2) assms(3) leq_mix_imp_weak_inferior my_U_represents_pref by blast qed lemma my_U_is_linear_function: assumes "p \ \

" and "q \ \

" and "\ \ {0..1}" assumes "some_best \[\] some_worst" shows "my_U (mix_pmf \ p q) = \ * my_U p + (1 - \) * my_U q" proof - define B where B: "B = some_best" define W where W:"W = some_worst" define Up where Up: "Up = my_U p" define Uq where Uq: "Uq = my_U q" have long_in: "(\ * Up + (1 - \) * Uq) \ {0..1}" proof - have "Up \ {0..1}" using assms Up my_U_is_defined(1) by blast moreover have "Uq \ {0..1}" using assms Uq my_U_is_defined(1) by blast moreover have "\ * Up \ {0..1}" using \Up \ {0..1}\ assms(3) mult_le_one by auto moreover have "1-\ \ {0..1}" using assms(3) by auto moreover have "(1 - \) * Uq \ {0..1}" using mult_le_one[of "1-\" Uq] calculation(2) calculation(4) by auto ultimately show ?thesis using add_nonneg_nonneg[of "\ * Up" "(1 - \) * Uq"] convex_bound_le[of Up 1 Uq \ "1-\"] by simp qed have fst: "p \[\] (mix_pmf Up B W)" using assms my_U_is_defined[of p] B W Up by simp have snd: "q \[\] (mix_pmf Uq B W)" using assms my_U_is_defined[of q] B W Uq by simp have mp_in: "(mix_pmf Up B W) \ \

" using fst relation_in_carrier by blast have f2: "mix_pmf \ p q \[\] mix_pmf \ (mix_pmf Up B W) q" using fst assms(2) assms(3) mix_pmf_preferred_independence by blast have **: "mix_pmf \ (mix_pmf Up B W) (mix_pmf Uq B W) = mix_pmf (\ * Up + (1-\) * Uq) B W" (is "?L = ?R") proof - let ?mixPQ = "(mix_pmf (\ * Up + (1 - \) * Uq) B W)" have "\e\set_pmf ?L. pmf (?L) e = pmf ?mixPQ e" proof fix e assume asm: "e \ set_pmf ?L" have i1: "pmf (?L) e = \ * pmf (mix_pmf Up B W) e + pmf (mix_pmf Uq B W) e - \ * pmf (mix_pmf Uq B W) e" using pmf_mix_deeper[of \ "mix_pmf Up B W" "(mix_pmf Uq B W)" e] assms(3) by blast have i3: "... = \ * Up * pmf B e + \ * pmf W e - \ * Up * pmf W e + Uq * pmf B e + pmf W e - Uq * pmf W e - \ * Uq * pmf B e - \ * pmf W e + \ * Uq * pmf W e" using left_diff_distrib' pmf_mix_deeper[of Up B W e] pmf_mix_deeper[of Uq B W e] assms Up Uq my_U_is_defined(1) by (simp add: distrib_left right_diff_distrib) have j4: "pmf ?mixPQ e = (\ * Up + (1 - \) * Uq) * pmf B e + pmf W e - (\ * Up + (1 - \) * Uq) * pmf W e" using pmf_mix_deeper[of "(\ * Up + (1 - \) * Uq)" B W e] long_in by blast then show "pmf (?L) e = pmf ?mixPQ e" by (simp add: i1 i3 mult.commute right_diff_distrib' ring_class.ring_distribs(1)) qed then show ?thesis using pmf_equiv_intro1 by blast qed have "mix_pmf \ (mix_pmf Up B W) q \[\] ?L" using approx_remains_after_same_comp_left assms(3) mp_in snd by blast hence *: "mix_pmf \ p q \[\] mix_pmf \ (mix_pmf (my_U p) B W) (mix_pmf (my_U q) B W)" using Up Uq f2 trans_approx by blast have "mix_pmf \ (mix_pmf (my_U p) B W) (mix_pmf (my_U q) B W) = ?R" using Up Uq ** by blast hence "my_U (mix_pmf \ p q) = \ * Up + (1-\) * Uq" by (metis * B W assms(4) indiff_imp_same_utility_value long_in my_U_is_defined(1) my_U_is_defined(2) my_U_represents_pref relation_in_carrier) then show ?thesis using Up Uq by blast qed text \ Now we define a more general Utility function that also takes the degenerate case into account \ private definition general_U where "general_U p = (if some_best \[\] some_worst then 1 else my_U p)" lemma general_U_is_linear_function: assumes "p \ \

" and "q \ \

" and "\ \ {0..1}" shows "general_U (mix_pmf \ p q) = \ * (general_U p) + (1 - \) * (general_U q)" proof - consider "some_best \[\] some_worst" | "some_best \[\] some_worst" using best_def some_best_in_best some_worst_in_worst worst_def by auto then show ?thesis proof (cases, goal_cases) case 1 then show ?case using assms(1) assms(2) assms(3) general_U_def my_U_is_linear_function by auto next case 2 then show ?case using assms(1) assms(2) assms(3) general_U_def by auto qed qed lemma general_U_ordinal_Utility: shows "ordinal_utility \

\ general_U" proof (standard, goal_cases) case (1 x y) consider (a) "some_best \[\] some_worst" | (b) "some_best \[\] some_worst" using best_def some_best_in_best some_worst_in_worst worst_def by auto then show ?case proof (cases, goal_cases) case a have "some_best \[\] some_worst" using a by auto then show "x \[\] y = (general_U y \ general_U x)" using 1 my_U_represents_pref[of x y] general_U_def by simp next case b have "general_U x = 1" "general_U y = 1" by (simp add: b general_U_def)+ moreover have "x \[\] y" using b by (meson "1"(1) "1"(2) best_worst_indiff_all_indiff(1) some_best_in_best some_worst_in_worst trans_approx) ultimately show "x \[\] y = (general_U y \ general_U x)" using general_U_def by linarith qed next case (2 x y) then show ?case using relation_in_carrier by blast next case (3 x y) then show ?case using relation_in_carrier by blast qed text \ Proof of the linearity of general-U. If we consider the definition of expected utility functions from Maschler, Solan, Zamir we are done. \ theorem is_linear: assumes "p \ \

" and "q \ \

" and "\ \ {0..1}" shows "\u. u (mix_pmf \ p q) = \ * (u p) + (1-\) * (u q)" proof let ?u = "general_U" consider "some_best \[\] some_worst" | "some_best \[\] some_worst" using best_def some_best_in_best some_worst_in_worst worst_def by auto then show "?u (mix_pmf \ p q) = \ * ?u p + (1 - \) * ?u q" proof (cases) case 1 then show ?thesis using assms(1) assms(2) assms(3) general_U_def my_U_is_linear_function by auto next case 2 then show ?thesis by (simp add: general_U_def) qed qed text \ Now I define a Utility function that assigns a utility to all outcomes. These are only finitely many \ private definition ocU where "ocU p = general_U (return_pmf p)" lemma geral_U_is_expected_value_of_ocU: assumes "set_pmf p \ outcomes" shows "general_U p = measure_pmf.expectation p ocU" using fnt assms proof (induct rule: pmf_mix_induct') case (mix p q a) hence "general_U (mix_pmf a p q) = a * general_U p + (1-a) * general_U q" using general_U_is_linear_function[of p q a] mix.hyps assms lotteries_on_def mix.hyps by auto also have "... = a * measure_pmf.expectation p ocU + (1-a) * measure_pmf.expectation q ocU" by (simp add: mix.hyps(4) mix.hyps(5)) also have "... = measure_pmf.expectation (mix_pmf a p q) ocU" using general_U_is_linear_function expected_value_mix_pmf_distrib fnt infinite_super mix.hyps(1) by (metis fnt mix.hyps(2) mix.hyps(3)) finally show ?case . qed (auto simp: support_in_outcomes assms fnt integral_measure_pmf_real ocU_def) lemma ordinal_utility_expected_value: "ordinal_utility \

\ (\x. measure_pmf.expectation x ocU)" proof (standard, goal_cases) case (1 x y) have ocs: "set_pmf x \ outcomes" "set_pmf y \ outcomes" by (meson "1" subsetI support_in_outcomes)+ have "x \[\] y \ (measure_pmf.expectation y ocU \ measure_pmf.expectation x ocU)" proof - assume "x \[\] y" have "general_U x \ general_U y" by (meson \x \[\] y\ general_U_ordinal_Utility ordinal_utility_def) then show "(measure_pmf.expectation y ocU \ measure_pmf.expectation x ocU)" using geral_U_is_expected_value_of_ocU ocs by auto qed moreover have "(measure_pmf.expectation y ocU \ measure_pmf.expectation x ocU) \ x \[\] y" proof - assume "(measure_pmf.expectation y ocU \ measure_pmf.expectation x ocU)" then have "general_U x \ general_U y" by (simp add: geral_U_is_expected_value_of_ocU ocs(1) ocs(2)) then show "x \[\] y" by (meson "1"(1) "1"(2) general_U_ordinal_Utility ordinal_utility.util_def) qed ultimately show ?case by blast next case (2 x y) then show ?case using relation_in_carrier by blast next case (3 x y) then show ?case using relation_in_carrier by auto qed lemma ordinal_utility_expected_value': "\u. ordinal_utility \

\ (\x. measure_pmf.expectation x u)" using ordinal_utility_expected_value by blast lemma ocU_is_expected_utility_bernoulli: shows "\x \ \

. \y \ \

. x \[\] y \ measure_pmf.expectation x ocU \ measure_pmf.expectation y ocU" using ordinal_utility_expected_value by (meson ordinal_utility.util_def) end (* continuous *) end(* finite outcomes *) end (* system U *) lemma expected_value_is_utility_function: assumes fnt: "finite outcomes" and "outcomes \ {}" assumes "x \ lotteries_on outcomes" and "y \ lotteries_on outcomes" assumes "ordinal_utility (lotteries_on outcomes) \ (\x. measure_pmf.expectation x u)" shows "measure_pmf.expectation x u \ measure_pmf.expectation y u \ x \[\] y" (is "?L \ ?R") using assms(3) assms(4) assms(5) ordinal_utility.util_def_conf ordinal_utility.ordinal_utility_left iffI by (metis (no_types, lifting)) lemma system_U_implies_vNM_utility: assumes fnt: "finite outcomes" and "outcomes \ {}" assumes rpr: "rational_preference (lotteries_on outcomes) \" assumes ind: "independent_vnm (lotteries_on outcomes) \" assumes cnt: "continuous_vnm (lotteries_on outcomes) \" shows "\u. ordinal_utility (lotteries_on outcomes) \ (\x. measure_pmf.expectation x u)" using ordinal_utility_expected_value'[of outcomes \] assms by blast lemma vNM_utility_implies_rationality: assumes fnt: "finite outcomes" and "outcomes \ {}" assumes "\u. ordinal_utility (lotteries_on outcomes) \ (\x. measure_pmf.expectation x u)" shows "rational_preference (lotteries_on outcomes) \" using assms(3) ordinal_util_imp_rat_prefs by blast theorem vNM_utility_implies_independence: assumes fnt: "finite outcomes" and "outcomes \ {}" assumes "\u. ordinal_utility (lotteries_on outcomes) \ (\x. measure_pmf.expectation x u)" shows "independent_vnm (lotteries_on outcomes) \" proof (rule independent_vnmI2) fix p q r and \::real assume a1: "p \ \

outcomes" assume a2: "q \ \

outcomes" assume a3: "r \ \

outcomes" assume a4: "\ \ {0<..1}" have in_lots: "mix_pmf \ p r \ lotteries_on outcomes" "mix_pmf \ q r \ lotteries_on outcomes" using a1 a3 a4 mix_in_lot apply fastforce using a2 a3 a4 mix_in_lot by fastforce have fnts: "finite (set_pmf p)" "finite (set_pmf q)" "finite (set_pmf r)" using a1 a2 a3 fnt infinite_super lotteries_on_def by blast+ obtain u where u: "ordinal_utility (lotteries_on outcomes) \ (\x. measure_pmf.expectation x u)" using assms by blast have "p \[\] q \ mix_pmf \ p r \[\] mix_pmf \ q r" proof - assume "p \[\] q" hence f: "measure_pmf.expectation p u \ measure_pmf.expectation q u" using u a1 a2 ordinal_utility.util_def by fastforce have "measure_pmf.expectation (mix_pmf \ p r) u \ measure_pmf.expectation (mix_pmf \ q r) u" proof - have "measure_pmf.expectation (mix_pmf \ p r) u = \ * measure_pmf.expectation p u + (1 - \) * measure_pmf.expectation r u" using expected_value_mix_pmf_distrib[of p r \ u] assms fnts a4 by fastforce moreover have "measure_pmf.expectation (mix_pmf \ q r) u = \ * measure_pmf.expectation q u + (1 - \) * measure_pmf.expectation r u" using expected_value_mix_pmf_distrib[of q r \ u] assms fnts a4 by fastforce ultimately show ?thesis using f using a4 by auto qed then show "mix_pmf \ p r \[\] mix_pmf \ q r" using u ordinal_utility_expected_value' ocU_is_expected_utility_bernoulli in_lots by (simp add: in_lots ordinal_utility_def) qed moreover have "mix_pmf \ p r \[\] mix_pmf \ q r \ p \[\] q" proof - assume "mix_pmf \ p r \[\] mix_pmf \ q r" hence f:"measure_pmf.expectation (mix_pmf \ p r) u \ measure_pmf.expectation (mix_pmf \ q r) u" using ordinal_utility.ordinal_utility_left u by fastforce hence "measure_pmf.expectation p u \ measure_pmf.expectation q u" proof - have "measure_pmf.expectation (mix_pmf \ p r) u = \ * measure_pmf.expectation p u + (1 - \) * measure_pmf.expectation r u" using expected_value_mix_pmf_distrib[of p r \ u] assms fnts a4 by fastforce moreover have "measure_pmf.expectation (mix_pmf \ q r) u = \ * measure_pmf.expectation q u + (1 - \) * measure_pmf.expectation r u" using expected_value_mix_pmf_distrib[of q r \ u] assms fnts a4 by fastforce ultimately show ?thesis using f using a4 by auto qed then show "p \[\] q" using a1 a2 ordinal_utility.util_def_conf u by fastforce qed ultimately show "p \[\] q = mix_pmf \ p r \[\] mix_pmf \ q r" by blast qed lemma exists_weight_for_equality: assumes "a > c" and "a \ b" and "b \ c" shows "\(e::real) \ {0..1}. (1-e) * a + e * c = b" proof - from assms have "b \ closed_segment a c" by (simp add: closed_segment_eq_real_ivl) thus ?thesis by (auto simp: closed_segment_def) qed lemma vNM_utilty_implies_continuity: assumes fnt: "finite outcomes" and "outcomes \ {}" assumes "\u. ordinal_utility (lotteries_on outcomes) \ (\x. measure_pmf.expectation x u)" shows "continuous_vnm (lotteries_on outcomes) \" proof (rule continuous_vnmI) fix p q r assume a1: "p \ \

outcomes" assume a2: "q \ \

outcomes" assume a3: "r \ \

outcomes " assume a4: "p \[\] q \ q \[\] r" then have g: "p \[\] r" by (meson assms(3) ordinal_utility.util_imp_trans transD) obtain u where u: "ordinal_utility (lotteries_on outcomes) \ (\x. measure_pmf.expectation x u)" using assms by blast have geqa: "measure_pmf.expectation p u \ measure_pmf.expectation q u" "measure_pmf.expectation q u \ measure_pmf.expectation r u" using a4 u by (meson ordinal_utility.ordinal_utility_left)+ have fnts: "finite p" "finite q" "finite r" using a1 a2 a3 fnt infinite_super lotteries_on_def by auto+ consider "p \[\] r" | "p \[\] r" using g by auto then show "\\\{0..1}. mix_pmf \ p r \[\] q" proof (cases) case 1 define a where a: "a = measure_pmf.expectation p u" define b where b: "b = measure_pmf.expectation r u" define c where c: "c = measure_pmf.expectation q u" have "a > b" using "1" a1 a2 a3 a b ordinal_utility.util_def_conf u by force have "c \ a" "b \ c" using geqa a b c by blast+ then obtain e ::real where e: "e \ {0..1}" "(1-e) * a + e * b = c" using exists_weight_for_equality[of b a c] \b < a\ by blast have *:"1-e \ {0..1}" using e(1) by auto hence "measure_pmf.expectation (mix_pmf (1-e) p r) u = (1-e) * measure_pmf.expectation p u + e * measure_pmf.expectation r u" using expected_value_mix_pmf_distrib[of p r "1-e" u] fnts by fastforce also have "... = (1-e) * a + e * b" using a b by auto also have "... = c" using c e by auto finally have f: "measure_pmf.expectation (mix_pmf (1-e) p r) u = measure_pmf.expectation q u" using c by blast hence "mix_pmf (1-e) p r \[\] q" using expected_value_is_utility_function[of outcomes "mix_pmf (1-e) p r" q \ u] * proof - have "mix_pmf (1 - e) p r \ \

outcomes" using \1 - e \ {0..1}\ a1 a3 mix_in_lot by blast then show ?thesis using f a2 ordinal_utility.util_def u by fastforce qed then show ?thesis using exists_weight_for_equality expected_value_mix_pmf_distrib * by blast next case 2 have "r \[\] q" by (meson "2" a4 assms(3) ordinal_utility.util_imp_trans transD) then show ?thesis by force qed qed theorem Von_Neumann_Morgenstern_Utility_Theorem: assumes fnt: "finite outcomes" and "outcomes \ {}" shows "rational_preference (lotteries_on outcomes) \ \ independent_vnm (lotteries_on outcomes) \ \ continuous_vnm (lotteries_on outcomes) \ \ (\u. ordinal_utility (lotteries_on outcomes) \ (\x. measure_pmf.expectation x u))" using vNM_utility_implies_independence[OF assms, of \] system_U_implies_vNM_utility[OF assms, of \] vNM_utilty_implies_continuity[OF assms, of \] ordinal_util_imp_rat_prefs[of "lotteries_on outcomes" \] by auto end diff --git a/thys/Probabilistic_Prime_Tests/Algebraic_Auxiliaries.thy b/thys/Probabilistic_Prime_Tests/Algebraic_Auxiliaries.thy --- a/thys/Probabilistic_Prime_Tests/Algebraic_Auxiliaries.thy +++ b/thys/Probabilistic_Prime_Tests/Algebraic_Auxiliaries.thy @@ -1,480 +1,480 @@ (* File: Algebraic_Auxiliaries.thy Authors: Daniel Stüwe Miscellaneous facts about algebra and number theory *) section \Auxiliary Material\ theory Algebraic_Auxiliaries imports "HOL-Algebra.Algebra" "HOL-Computational_Algebra.Squarefree" "HOL-Number_Theory.Number_Theory" begin hide_const (open) Divisibility.prime lemma sum_of_bool_eq_card: assumes "finite S" shows "(\a \ S. of_bool (P a)) = real (card {a \ S . P a })" proof - have "(\a \ S. of_bool (P a) :: real) = (\a \ {x\S. P x}. 1)" using assms by (intro sum.mono_neutral_cong_right) auto thus ?thesis by simp qed lemma mod_natE: fixes a n b :: nat assumes "a mod n = b" shows "\ l. a = n * l + b" using assms mod_mult_div_eq[of a n] by (metis add.commute) lemma (in group) r_coset_is_image: "H #> a = (\ x. x \ a) ` H" unfolding r_coset_def image_def by blast lemma (in group) FactGroup_order: assumes "subgroup H G" "finite H" shows "order G = order (G Mod H) * card H" using lagrange assms unfolding FactGroup_def order_def by simp corollary (in group) FactGroup_order_div: assumes "subgroup H G" "finite H" shows "order (G Mod H) = order G div card H" using assms FactGroup_order subgroupE(2)[OF \subgroup H G\] by (auto simp: order_def) lemma group_hom_imp_group_hom_image: assumes "group_hom G G h" shows "group_hom G (G\carrier := h ` carrier G\) h" using group_hom.axioms[OF assms] group_hom.img_is_subgroup[OF assms] group.subgroup_imp_group by(auto intro!: group_hom.intro simp: group_hom_axioms_def hom_def) theorem homomorphism_thm: assumes "group_hom G G h" shows "G Mod kernel G (G\carrier := h ` carrier G\) h \ G \carrier := h ` carrier G\" by (intro group_hom.FactGroup_iso group_hom_imp_group_hom_image assms) simp lemma is_iso_imp_same_card: assumes "H \ G " shows "order H = order G" proof - from assms obtain h where "bij_betw h (carrier H) (carrier G)" unfolding is_iso_def iso_def by blast then show ?thesis unfolding order_def by (rule bij_betw_same_card) qed corollary homomorphism_thm_order: assumes "group_hom G G h" shows "order (G\carrier := h ` carrier G\) * card (kernel G (G\carrier := h ` carrier G\) h) = order G " proof - have "order (G\carrier := h ` carrier G\) = order (G Mod (kernel G (G\carrier := h ` carrier G\) h))" using is_iso_imp_same_card[OF homomorphism_thm] \group_hom G G h\ by fastforce moreover have "group G" using \group_hom G G h\ group_hom.axioms by blast ultimately show ?thesis using \group_hom G G h\ and group_hom_imp_group_hom_image[OF \group_hom G G h\] unfolding FactGroup_def by (simp add: group.lagrange group_hom.subgroup_kernel order_def) qed lemma (in group_hom) kernel_subset: "kernel G H h \ carrier G" using subgroup_kernel G.subgroupE(1) by blast lemma (in group) proper_subgroup_imp_bound_on_card: assumes "H \ carrier G" "subgroup H G" "finite (carrier G)" shows "card H \ order G div 2" proof - from \finite (carrier G)\ have "finite (rcosets H)" by (simp add: RCOSETS_def) note subgroup.subgroup_in_rcosets[OF \subgroup H G\ is_group] then obtain J where "J \ H" "J \ rcosets H" using rcosets_part_G[OF \subgroup H G\] and \H \ carrier G\ by (metis Sup_le_iff inf.absorb_iff2 inf.idem inf.strict_order_iff) then have "2 \ card (rcosets H)" using \H \ rcosets H\ card_mono[OF \finite (rcosets H)\, of "{H, J}"] by simp then show ?thesis using mult_le_mono[of 2 "card (rcosets H)" "card H" "card H"] unfolding lagrange[OF \subgroup H G\] by force qed lemma cong_exp_trans[trans]: "[a ^ b = c] (mod n) \ [a = d] (mod n) \ [d ^ b = c] (mod n)" "[c = a ^ b] (mod n) \ [a = d] (mod n) \ [c = d ^ b] (mod n)" using cong_pow cong_sym cong_trans by blast+ lemma cong_exp_mod[simp]: "[(a mod n) ^ b = c] (mod n) \ [a ^ b = c] (mod n)" "[c = (a mod n) ^ b] (mod n) \ [c = a ^ b] (mod n)" by (auto simp add: cong_def mod_simps) lemma cong_mult_mod[simp]: "[(a mod n) * b = c] (mod n) \ [a * b = c] (mod n)" "[a * (b mod n) = c] (mod n) \ [a * b = c] (mod n)" by (auto simp add: cong_def mod_simps) lemma cong_add_mod[simp]: "[(a mod n) + b = c] (mod n) \ [a + b = c] (mod n)" "[a + (b mod n) = c] (mod n) \ [a + b = c] (mod n)" "[\i\A. f i mod n = c] (mod n) \ [\i\A. f i = c] (mod n)" by (auto simp add: cong_def mod_simps) lemma cong_add_trans[trans]: "[a = b + x] (mod n) \ [x = y] (mod n) \ [a = b + y] (mod n)" "[a = x + b] (mod n) \ [x = y] (mod n) \ [a = y + b] (mod n)" "[b + x = a] (mod n) \ [x = y] (mod n) \ [b + y = a] (mod n)" "[x + b = a] (mod n) \ [x = y] (mod n) \ [y + b = a] (mod n)" unfolding cong_def using mod_simps(1, 2) by metis+ lemma cong_mult_trans[trans]: "[a = b * x] (mod n) \ [x = y] (mod n) \ [a = b * y] (mod n)" "[a = x * b] (mod n) \ [x = y] (mod n) \ [a = y * b] (mod n)" "[b * x = a] (mod n) \ [x = y] (mod n) \ [b * y = a] (mod n)" "[x * b = a] (mod n) \ [x = y] (mod n) \ [y * b = a] (mod n)" unfolding cong_def using mod_simps(4, 5) by metis+ lemma cong_diff_trans[trans]: "[a = b - x] (mod n) \ [x = y] (mod n) \ [a = b - y] (mod n)" "[a = x - b] (mod n) \ [x = y] (mod n) \ [a = y - b] (mod n)" "[b - x = a] (mod n) \ [x = y] (mod n) \ [b - y = a] (mod n)" "[x - b = a] (mod n) \ [x = y] (mod n) \ [y - b = a] (mod n)" for a :: "'a :: {unique_euclidean_semiring, euclidean_ring_cancel}" unfolding cong_def by (metis mod_diff_eq)+ lemma eq_imp_eq_mod_int: "a = b \ [a = b] (mod m)" for a b :: int by simp lemma eq_imp_eq_mod_nat: "a = b \ [a = b] (mod m)" for a b :: nat by simp lemma cong_pow_I: "a = b \ [x^a = x^b](mod n)" by simp lemma gre1I: "(n = 0 \ False) \ (1 :: nat) \ n" by presburger lemma gre1I_nat: "(n = 0 \ False) \ (Suc 0 :: nat) \ n" by presburger lemma totient_less_not_prime: assumes "\ prime n" "1 < n" shows "totient n < n - 1" using totient_imp_prime totient_less assms by (metis One_nat_def Suc_pred le_less_trans less_SucE zero_le_one) lemma power2_diff_nat: "x \ y \ (x - y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 - 2 * x * y" for x y :: nat by (simp add: algebra_simps power2_eq_square mult_2_right) (meson Nat.diff_diff_right le_add2 le_trans mult_le_mono order_refl) lemma square_inequality: "1 < n \ (n + n) \ (n * n)" for n :: nat by (metis Suc_eq_plus1_left Suc_leI mult_le_mono1 semiring_normalization_rules(4)) lemma square_one_cong_one: assumes "[x = 1](mod n)" shows "[x^2 = 1](mod n)" using assms cong_pow by fastforce lemma cong_square_alt_int: "prime p \ [a * a = 1] (mod p) \ [a = 1] (mod p) \ [a = p - 1] (mod p)" for a p :: "'a :: {normalization_semidom, linordered_idom, unique_euclidean_ring}" using dvd_add_triv_right_iff[of p "a - (p - 1)"] by (auto simp add: cong_iff_dvd_diff square_diff_one_factored dest!: prime_dvd_multD) lemma cong_square_alt: "prime p \ [a * a = 1] (mod p) \ [a = 1] (mod p) \ [a = p - 1] (mod p)" for a p :: nat - using cong_square_alt_int and cong_int_iff prime_nat_int_transfer - by (metis (mono_tags) int_ops(2) int_ops(7) less_imp_le_nat of_nat_diff prime_gt_1_nat) + using cong_square_alt_int[of "int p" "int a"] prime_nat_int_transfer[of p] prime_gt_1_nat[of p] + by (simp flip: cong_int_iff add: of_nat_diff) lemma square_minus_one_cong_one: fixes n x :: nat assumes "1 < n" "[x = n - 1](mod n)" shows "[x^2 = 1](mod n)" proof - have "[x^2 = (n - 1) * (n - 1)] (mod n)" using cong_mult[OF assms(2) assms(2)] by (simp add: algebra_simps power2_eq_square) also have "[(n - 1) * (n - 1) = Suc (n * n) - (n + n)] (mod n)" using power2_diff_nat[of 1 n] \1 < n\ by (simp add: algebra_simps power2_eq_square) also have "[Suc (n * n) - (n + n) = Suc (n * n)] (mod n)" proof - have "n * n + 0 * n = n * n" by linarith moreover have "n * n - (n + n) + (n + n) = n * n" using square_inequality[OF \1 < n\] le_add_diff_inverse2 by blast moreover have "(Suc 0 + 1) * n = n + n" by simp ultimately show ?thesis using square_inequality[OF \1 < n\] by (metis (no_types) Suc_diff_le add_Suc cong_iff_lin_nat) qed also have "[Suc (n * n) = 1] (mod n)" using cong_to_1'_nat by auto finally show ?thesis . qed lemma odd_prime_gt_2_int: "2 < p" if "odd p" "prime p" for p :: int using prime_ge_2_int[OF \prime p\] \odd p\ by (cases "p = 2") auto lemma odd_prime_gt_2_nat: "2 < p" if "odd p" "prime p" for p :: nat using prime_ge_2_nat[OF \prime p\] \odd p\ by (cases "p = 2") auto lemma gt_one_imp_gt_one_power_if_coprime: "1 \ x \ 1 < n \ coprime x n \ 1 \ x ^ (n - 1) mod n" by (rule gre1I) (auto simp: coprime_commute dest: coprime_absorb_left) lemma residue_one_dvd: "a mod n = 1 \ n dvd a - 1" for a n :: nat by (fastforce intro!: cong_to_1_nat simp: cong_def) lemma coprimeI_power_mod: fixes x r n :: nat assumes "x ^ r mod n = 1" "r \ 0" "n \ 0" shows "coprime x n" proof - have "coprime (x ^ r mod n) n" using coprime_1_right \x ^ r mod n = 1\ by (simp add: coprime_commute) thus ?thesis using \r \ 0\ \n \ 0\ by simp qed (* MOVE - EXTRA *) lemma prime_dvd_choose: assumes "0 < k" "k < p" "prime p" shows "p dvd (p choose k)" proof - have "k \ p" using \k < p\ by auto have "p dvd fact p" using \prime p\ by (simp add: prime_dvd_fact_iff) moreover have "\ p dvd fact k * fact (p - k)" unfolding prime_dvd_mult_iff[OF \prime p\] prime_dvd_fact_iff[OF \prime p\] using assms by simp ultimately show ?thesis unfolding binomial_fact_lemma[OF \k \ p\, symmetric] using assms prime_dvd_multD by blast qed lemma cong_eq_0_I: "(\i\A. [f i mod n = 0] (mod n)) \ [\i\A. f i = 0] (mod n)" using cong_sum by fastforce lemma power_mult_cong: assumes "[x^n = a](mod m)" "[y^n = b](mod m)" shows "[(x*y)^n = a*b](mod m)" using assms cong_mult[of "x^n" a m "y^n" b] power_mult_distrib by metis lemma fixes n :: nat assumes "n > 1" shows odd_pow_cong: "odd m \ [(n - 1) ^ m = n - 1] (mod n)" and even_pow_cong: "even m \ [(n - 1) ^ m = 1] (mod n)" proof (induction m) case (Suc m) case 1 with Suc have IH: "[(n - 1) ^ m = 1] (mod n)" by auto show ?case using \1 < n\ cong_mult[OF cong_refl IH] by simp next case (Suc m) case 2 with Suc have IH: "[(n - 1) ^ m = n - 1] (mod n)" by auto show ?case using cong_mult[OF cong_refl IH, of "(n - 1)"] and square_minus_one_cong_one[OF \1 < n\, of "n - 1"] by (auto simp: power2_eq_square intro: cong_trans) qed simp_all lemma cong_mult_uneq': fixes a :: "'a::{unique_euclidean_ring, ring_gcd}" assumes "coprime d a" shows "[b \ c] (mod a) \ [d = e] (mod a) \ [b * d \ c * e] (mod a)" using cong_mult_rcancel[OF assms] using cong_trans[of "b*d" "c*e" a "c*d"] using cong_scalar_left cong_sym by blast lemma p_coprime_right_nat: "prime p \ coprime a p = (\ p dvd a)" for p a :: nat by (meson coprime_absorb_left coprime_commute not_prime_unit prime_imp_coprime_nat) lemma squarefree_mult_imp_coprime: assumes "squarefree (a * b :: 'a :: semiring_gcd)" shows "coprime a b" proof (rule coprimeI) fix l assume "l dvd a" "l dvd b" then obtain a' b' where "a = l * a'" "b = l * b'" by (auto elim!: dvdE) with assms have "squarefree (l\<^sup>2 * (a' * b'))" by (simp add: power2_eq_square mult_ac) thus "l dvd 1" by (rule squarefreeD) auto qed lemma prime_divisor_exists_strong: fixes m :: int assumes "m > 1" "\prime m" shows "\n k. m = n * k \ 1 < n \ n < m \ 1 < k \ k < m" proof - from assms obtain n k where nk: "n * k > 1" "n \ 0" "m = n * k" "n \ 1" "n \ 0" "k \ 1" using assms unfolding prime_int_iff dvd_def by auto from nk have "n > 1" by linarith from nk assms have "n * k > 0" by simp with \n \ 0\ have "k > 0" using zero_less_mult_pos by force with \k \ 1\ have "k > 1" by linarith from nk have "n > 1" by linarith from \k > 1\ nk have "n < m" "k < m" by simp_all with nk \k > 1\ \n > 1\ show ?thesis by blast qed lemma prime_divisor_exists_strong_nat: fixes m :: nat assumes "1 < m" "\prime m" shows "\p k. m = p * k \ 1 < p \ p < m \ 1 < k \ k < m \ prime p" proof - obtain p where p_def: "prime p" "p dvd m" "p \ m" "1 < p" using assms prime_prime_factor and prime_gt_1_nat by blast moreover define k where "k = m div p" with \p dvd m\ have "m = p * k" by simp moreover have "p < m" using \p \ m\ dvd_imp_le[OF \p dvd m\] and \m > 1\ by simp moreover have "1 < k" "k < m" using \1 < m\ \1 < p\ and \p \ m\ unfolding \m = p * k\ by (force intro: Suc_lessI Nat.gr0I)+ ultimately show ?thesis using \1 < m\ by blast qed (* TODO Remove *) lemma prime_factorization_eqI: assumes "\p. p \# P \ prime p" "prod_mset P = n" shows "prime_factorization n = P" using prime_factorization_prod_mset_primes[of P] assms by simp lemma prime_factorization_prime_elem: assumes "prime_elem p" shows "prime_factorization p = {#normalize p#}" proof - have "prime_factorization p = prime_factorization (normalize p)" by (metis normalize_idem prime_factorization_cong) also have "\ = {#normalize p#}" by (rule prime_factorization_prime) (use assms in auto) finally show ?thesis . qed lemma size_prime_factorization_eq_Suc_0_iff [simp]: fixes n :: "'a :: factorial_semiring_multiplicative" shows "size (prime_factorization n) = Suc 0 \ prime_elem n" proof assume size: "size (prime_factorization n) = Suc 0" hence [simp]: "n \ 0" by auto from size obtain p where *: "prime_factorization n = {#p#}" by (auto elim!: size_mset_SucE) hence p: "p \ prime_factors n" by auto have "prime_elem (normalize p)" using p by (auto simp: in_prime_factors_iff) also have "p = prod_mset (prime_factorization n)" using * by simp also have "normalize \ = normalize n" by (rule prod_mset_prime_factorization_weak) auto finally show "prime_elem n" by simp qed (auto simp: prime_factorization_prime_elem) (* END TODO *) (* TODO Move *) lemma squarefree_prime_elem [simp, intro]: fixes p :: "'a :: algebraic_semidom" assumes "prime_elem p" shows "squarefree p" proof (rule squarefreeI) fix x assume "x\<^sup>2 dvd p" show "is_unit x" proof (rule ccontr) assume "\is_unit x" hence "\is_unit (x\<^sup>2)" by (simp add: is_unit_power_iff) from assms and this and \x\<^sup>2 dvd p\ have "prime_elem (x\<^sup>2)" by (rule prime_elem_mono) thus False by (simp add: prime_elem_power_iff) qed qed lemma squarefree_prime [simp, intro]: "prime p \ squarefree p" by auto lemma not_squarefree_primepow: assumes "primepow n" shows "squarefree n \ prime n" using assms by (auto simp: primepow_def squarefree_power_iff prime_power_iff) lemma prime_factorization_normalize [simp]: "prime_factorization (normalize n) = prime_factorization n" by (rule prime_factorization_cong) auto lemma one_prime_factor_iff_primepow: fixes n :: "'a :: factorial_semiring_multiplicative" shows "card (prime_factors n) = Suc 0 \ primepow (normalize n)" proof assume "primepow (normalize n)" then obtain p k where pk: "prime p" "normalize n = p ^ k" "k > 0" by (auto simp: primepow_def) hence "card (prime_factors (normalize n)) = Suc 0" by (subst pk) (simp add: prime_factors_power prime_factorization_prime) thus "card (prime_factors n) = Suc 0" by simp next assume *: "card (prime_factors n) = Suc 0" from * have "(\p\prime_factors n. p ^ multiplicity p n) = normalize n" by (intro prod_prime_factors) auto also from * have "card (prime_factors n) = 1" by simp then obtain p where p: "prime_factors n = {p}" by (elim card_1_singletonE) finally have "normalize n = p ^ multiplicity p n" by simp moreover from p have "prime p" "multiplicity p n > 0" by (auto simp: prime_factors_multiplicity) ultimately show "primepow (normalize n)" unfolding primepow_def by blast qed lemma squarefree_imp_prod_prime_factors_eq: fixes x :: "'a :: factorial_semiring_multiplicative" assumes "squarefree x" shows "\(prime_factors x) = normalize x" proof - from assms have [simp]: "x \ 0" by auto have "(\p\prime_factors x. p ^ multiplicity p x) = normalize x" by (intro prod_prime_factors) auto also have "(\p\prime_factors x. p ^ multiplicity p x) = (\p\prime_factors x. p)" using assms by (intro prod.cong refl) (auto simp: squarefree_factorial_semiring') finally show ?thesis by simp qed (* END TODO *) end \ No newline at end of file diff --git a/thys/Stable_Matching/Basis.thy b/thys/Stable_Matching/Basis.thy --- a/thys/Stable_Matching/Basis.thy +++ b/thys/Stable_Matching/Basis.thy @@ -1,572 +1,597 @@ (*<*) theory Basis imports Main "HOL-Library.While_Combinator" begin (*>*) section\ Preliminaries \ (*>*)(*<*) subsection\ HOL Detritus \ lemma Above_union: shows "x \ Above r (X \ Y) \ x \ Above r X \ x \ Above r Y" unfolding Above_def by blast lemma Above_Field: assumes "x \ Above r X" shows "x \ Field r" using assms unfolding Above_def by blast lemma AboveS_Field: assumes "x \ AboveS r X" shows "x \ Field r" using assms unfolding AboveS_def by blast lemma Above_Linear_singleton: assumes "x \ Field r" assumes "Linear_order r" shows "x \ Above r {x}" using assms unfolding Above_def order_on_defs by (force dest: refl_onD) lemma subseqs_set: assumes "y \ set (subseqs xs)" shows "set y \ set xs" using assms by (metis Pow_iff image_eqI subseqs_powset) primrec map_of_default :: "'v \ ('k \ 'v) list \ 'k \ 'v" where "map_of_default v0 [] k = v0" | "map_of_default v0 (kv # kvs) k = (if k = fst kv then snd kv else map_of_default v0 kvs k)" lemmas set_elem_equalityI = Set.equalityI[OF Set.subsetI Set.subsetI] lemmas total_onI = iffD2[OF total_on_def, rule_format] lemma partial_order_on_acyclic: assumes "partial_order_on A r" shows "acyclic (r - Id)" by (metis acyclic_irrefl assms irrefl_diff_Id partial_order_on_def preorder_on_def trancl_id trans_diff_Id) lemma finite_Linear_order_induct[consumes 3, case_names step]: assumes "Linear_order r" assumes "x \ Field r" assumes "finite r" assumes step: "\x. \x \ Field r; \y. y \ aboveS r x \ P y\ \ P x" shows "P x" using assms(2) proof(induct rule: wf_induct[of "r\ - Id"]) from assms(1,3) show "wf (r\ - Id)" using linear_order_on_well_order_on linear_order_on_converse unfolding well_order_on_def by blast next case (2 x) then show ?case by - (rule step; auto simp: aboveS_def intro: FieldI2) qed text\ We sometimes want a notion of monotonicity over some set. \ definition mono_on :: "'a::order set \ ('a \ 'b::order) \ bool" where "mono_on A f = (\x\A. \y\A. x \ y \ f x \ f y)" lemmas mono_onI = iffD2[OF mono_on_def, rule_format] lemmas mono_onD = iffD1[OF mono_on_def, rule_format] lemma mono_onE: "\mono_on A f; x \ A; y \ A; x \ y; f x \ f y \ thesis\ \ thesis" using mono_onD by blast lemma mono_on_mono: "mono_on UNIV = mono" by (clarsimp simp: mono_on_def mono_def fun_eq_iff) (*>*) subsection\ MaxR: maximum elements of linear orders \ text\ We generalize the existing @{const "max"} and @{const "Max"} functions to work on orders defined over sets. See \S\ref{sec:cf-linear} for choice-function related lemmas. \ locale MaxR = fixes r :: "'a::finite rel" assumes r_Linear_order: "Linear_order r" begin text\ The basic function chooses the largest of two elements: \ definition maxR :: "'a \ 'a \ 'a" where "maxR x y = (if (x, y) \ r then y else x)" (*<*) lemma maxR_domain: shows "{x, y} \ A \ maxR x y \ A" unfolding maxR_def by simp lemma maxR_range: shows "maxR x y \ {x, y}" unfolding maxR_def by simp lemma maxR_rangeD: "maxR x y \ x \ maxR x y = y" "maxR x y \ y \ maxR x y = x" unfolding maxR_def by auto lemma maxR_idem: shows "maxR x x = x" unfolding maxR_def by simp lemma maxR_absorb2: shows "(x, y) \ r \ maxR x y = y" unfolding maxR_def by simp lemma maxR_absorb1: shows "(y, x) \ r \ maxR x y = x" using r_Linear_order unfolding maxR_def by (simp add: order_on_defs antisym_def) lemma maxR_assoc: shows "{x,y,z} \ Field r \ maxR (maxR x y) z = maxR x (maxR y z)" using r_Linear_order unfolding maxR_def by simp (metis order_on_defs(1-3) total_on_def trans_def) lemma maxR_commute: shows "{x,y} \ Field r \ maxR x y = maxR y x" using r_Linear_order unfolding maxR_def by (fastforce simp: order_on_defs antisym_def total_on_def) lemmas maxR_simps = maxR_idem maxR_absorb1 maxR_absorb2 (*>*) text\ We hoist this to finite sets using the @{const "Finite_Set.fold"} combinator. For code generation purposes it seems inevitable that we need to fuse the fold and filter into a single total recursive definition. \ definition MaxR_f :: "'a \ 'a option \ 'a option" where "MaxR_f x acc = (if x \ Field r then Some (case acc of None \ x | Some y \ maxR x y) else acc)" interpretation MaxR_f: comp_fun_idem MaxR_f using %invisible r_Linear_order by unfold_locales (fastforce simp: fun_eq_iff maxR_def MaxR_f_def order_on_defs total_on_def antisymD elim: transE split: option.splits)+ definition MaxR_opt :: "'a set \ 'a option" where MaxR_opt_eq_fold': "MaxR_opt A = Finite_Set.fold MaxR_f None A" (*<*) lemma empty [simp]: shows "MaxR_opt {} = None" by (simp add: MaxR_opt_eq_fold') lemma shows insert: "MaxR_opt (insert x A) = (if x \ Field r then Some (case MaxR_opt A of None \ x | Some y \ maxR x y) else MaxR_opt A)" and range_Some[rule_format]: "MaxR_opt A = Some a \ a \ A \ Field r" using finite[of A] by induct (auto simp: MaxR_opt_eq_fold' maxR_def MaxR_f_def split: option.splits) lemma range_None: assumes "MaxR_opt A = None" shows "A \ Field r = {}" using assms by (metis Int_iff insert all_not_in_conv insert_absorb option.simps(3)) lemma domain_empty: assumes "A \ Field r = {}" shows "MaxR_opt A = None" using assms by (metis empty_iff option.exhaust range_Some) lemma domain: shows "MaxR_opt (A \ Field r) = MaxR_opt A" using finite[of A] by induct (simp_all add: insert) lemmas MaxR_opt_code = MaxR_opt_eq_fold'[where A="set A", unfolded MaxR_f.fold_set_fold] for A lemma range: shows "MaxR_opt A \ Some ` (A \ Field r) \ {None}" using range_Some notin_range_Some by fastforce lemma union: shows "MaxR_opt (A \ B) = (case MaxR_opt A of None \ MaxR_opt B | Some mA \ Some (case MaxR_opt B of None \ mA | Some mB \ maxR mA mB))" using finite[of A] by induct (auto simp: maxR_assoc insert dest!: range_Some split: option.splits) lemma mono: assumes "MaxR_opt A = Some x" shows "\y. MaxR_opt (A \ B) = Some y \ (x, y) \ r" using finite[of B] proof induct case empty with assms show ?case using range_Some underS_incl_iff[OF r_Linear_order] by fastforce next note ins = insert case (insert b B) with assms r_Linear_order show ?case unfolding order_on_defs total_on_def by (fastforce simp: ins maxR_def elim: transE intro: FieldI1) qed + +declare [[simproc del: eliminate_false_implies]] + lemma MaxR_opt_is_greatest: assumes "MaxR_opt A = Some x" assumes "y \ A \ Field r" shows "(y, x) \ r" using finite[of A] assms proof(induct arbitrary: x) note ins = insert - case (insert a A) then show ?case - using r_Linear_order unfolding order_on_defs refl_on_def total_on_def - by (auto 10 0 simp: maxR_def ins dest!: range_None range_Some split: if_splits option.splits elim: transE) + case (insert a A) + show ?case + proof (cases "y = x") + case True + thus "(y, x) \ r" + using r_Linear_order insert by (auto simp: order_on_defs refl_on_def) + next + case False + show "(y, x) \ r" + proof (rule ccontr) + assume "(y, x) \ r" + from insert have "x \ Field r" "y \ Field r" + by (auto simp: maxR_def ins dest!: range_None range_Some split: if_splits option.splits) + from \(y, x) \ r\ and \y \ x\ and insert obtain z where z: "(x, z) \ r" "(y, z) \ r" "z \ Field r" + by (auto simp: maxR_def ins dest!: range_None range_Some split: if_splits option.splits) + have "(x, y) \ r" + using r_Linear_order \(y, x) \ r\ \x \ Field r\ \y \ Field r\ \y \ x\ + by (auto simp: order_on_defs total_on_def) + have "trans r" + using r_Linear_order by (auto simp: order_on_defs) + from this and \(x, y) \ r\ and \(y, z) \ r\ have "(x, z) \ r" + by (rule transD) + with \(x, z) \ r\ show False by contradiction + qed + qed qed simp lemma greatest_is_MaxR_opt: assumes "x \ A \ Field r" assumes "\y \ A \ Field r. (y, x) \ r" shows "MaxR_opt A = Some x" using finite[of A] assms proof(induct arbitrary: x) note ins = insert case (insert a A) then show ?case using maxR_absorb1 maxR_absorb2 by (fastforce simp: maxR_def ins dest: range_None range_Some split: option.splits) qed simp lemma subset: assumes "set_option (MaxR_opt B) \ A" assumes "A \ B" shows "MaxR_opt B = MaxR_opt A" using union[where A=A and B="B-A"] range[of "B - A"] assms by (auto simp: Un_absorb1 finite_subset maxR_def split: option.splits) (*>*) end interpretation MaxR_empty: MaxR "{}" by unfold_locales simp interpretation MaxR_singleton: MaxR "{(x,x)}" for x by unfold_locales simp lemma MaxR_r_domain [iff]: assumes "MaxR r" shows "MaxR (Restr r A)" using assms Linear_order_Restr unfolding MaxR_def by blast subsection\ Linear orders from lists \ text\ Often the easiest way to specify a concrete linear order is with a list. Here these run from greatest to least. \ primrec linord_of_listP :: "'a \ 'a \ 'a list \ bool" where "linord_of_listP x y [] \ False" | "linord_of_listP x y (z # zs) \ (z = y \ x \ set (z # zs)) \ linord_of_listP x y zs" definition linord_of_list :: "'a list \ 'a rel" where "linord_of_list xs \ {(x, y). linord_of_listP x y xs}" (*<*) lemma linord_of_list_linord_of_listP: shows "xy \ linord_of_list xs \ linord_of_listP (fst xy) (snd xy) xs" unfolding linord_of_list_def split_def by simp lemma linord_of_listP_linord_of_list: shows "linord_of_listP x y xs \ (x, y) \ linord_of_list xs" unfolding linord_of_list_def by simp lemma linord_of_listP_empty: shows "(\x y. \linord_of_listP x y xs) \ xs = []" by (metis linord_of_listP.simps list.exhaust list.set_intros(1)) lemma linord_of_listP_domain: assumes "linord_of_listP x y xs" shows "x \ set xs \ y \ set xs" using assms by (induct xs) auto lemma linord_of_list_empty[iff]: "linord_of_list [] = {}" "linord_of_list xs = {} \ xs = []" unfolding linord_of_list_def by (simp_all add: linord_of_listP_empty) lemma linord_of_list_singleton: "(x, y) \ linord_of_list [z] \ x = z \ y = z" by (force simp: linord_of_list_linord_of_listP) lemma linord_of_list_range: "linord_of_list xs \ set xs \ set xs" unfolding linord_of_list_def by (induct xs) auto lemma linord_of_list_Field [simp]: "Field (linord_of_list xs) = set xs" unfolding linord_of_list_def by (induct xs) (auto simp: Field_def) lemma linord_of_listP_append: "linord_of_listP x y (xs @ ys) \ linord_of_listP x y xs \ linord_of_listP x y ys \ (y \ set xs \ x \ set ys)" by (induct xs) auto lemma linord_of_list_append: "(x, y) \ linord_of_list (xs @ ys) \ (x, y) \ linord_of_list xs \ (x, y) \ linord_of_list ys \ (y \ set xs \ x \ set ys)" unfolding linord_of_list_def by (simp add: linord_of_listP_append) lemma linord_of_list_refl_on: shows "refl_on (set xs) (linord_of_list xs)" unfolding linord_of_list_def by (induct xs) (auto intro!: refl_onI simp: refl_onD1 refl_onD2 dest: refl_onD subsetD[OF linord_of_list_range]) lemma linord_of_list_trans: assumes "distinct xs" shows "trans (linord_of_list xs)" using assms unfolding linord_of_list_def by (induct xs) (auto intro!: transI dest: linord_of_listP_domain elim: transE) lemma linord_of_list_antisym: assumes "distinct xs" shows "antisym (linord_of_list xs)" using assms unfolding linord_of_list_def by (induct xs) (auto intro!: antisymI dest: linord_of_listP_domain simp: antisymD) lemma linord_of_list_total_on: shows "total_on (set xs) (linord_of_list xs)" unfolding total_on_def linord_of_list_def by (induct xs) auto lemma linord_of_list_Restr: assumes "x \ C" notes in_set_remove1[simp del] (* suppress warning *) shows "Restr (linord_of_list (remove1 x xs)) C = Restr (linord_of_list xs) C" using assms unfolding linord_of_list_def by (induct xs) (auto iff: in_set_remove1) lemma linord_of_list_nth: assumes "(xs ! i, xs ! j) \ linord_of_list xs" assumes "i < length xs" "j < length xs" assumes "distinct xs" shows "j \ i" using %invisible assms proof(induct xs arbitrary: i j) case (Cons x xs i j) show ?case proof(cases "i < length xs") case True with Cons show ?thesis by (auto simp: linord_of_list_linord_of_listP nth_equal_first_eq less_Suc_eq_0_disj linord_of_listP_domain) next case False with Cons show ?thesis by fastforce qed qed simp (*>*) text\\ lemma linord_of_list_Linear_order: assumes "distinct xs" assumes "ys = set xs" shows "linear_order_on ys (linord_of_list xs)" using %invisible assms linord_of_list_range linord_of_list_refl_on linord_of_list_trans linord_of_list_antisym linord_of_list_total_on unfolding order_on_defs by force text\ Every finite linear order is generated by a list. \ (*<*) inductive sorted_on :: "'a rel \ 'a list \ bool" where Nil [iff]: "sorted_on r []" | Cons [intro!]: "\x \ Field r; \y\set xs. (x, y) \ r; sorted_on r xs\ \ sorted_on r (x # xs)" inductive_cases sorted_on_inv[elim!]: "sorted_on r []" "sorted_on r (x # xs)" primrec insort_key_on :: "'a rel \ ('b \ 'a) \ 'b \ 'b list \ 'b list" where "insort_key_on r f x [] = [x]" | "insort_key_on r f x (y # ys) = (if (f x, f y) \ r then (x # y # ys) else y # insort_key_on r f x ys)" definition sort_key_on :: "'a rel \ ('b \ 'a) \ 'b list \ 'b list" where "sort_key_on r f xs = foldr (insort_key_on r f) xs []" definition insort_insert_key_on :: "'a rel \ ('b \ 'a) \ 'b \ 'b list \ 'b list" where "insort_insert_key_on r f x xs = (if f x \ f ` set xs then xs else insort_key_on r f x xs)" abbreviation "sort_on r \ sort_key_on r (\x. x)" abbreviation "insort_on r \ insort_key_on r (\x. x)" abbreviation "insort_insert_on r \ insort_insert_key_on r (\x. x)" context fixes r :: "'a rel" assumes "Linear_order r" begin lemma sorted_on_single [iff]: shows "sorted_on r [x] \ x \ Field r" by (metis empty_iff list.distinct(1) list.set(1) nth_Cons_0 sorted_on.simps) lemma sorted_on_many: assumes "(x, y) \ r" assumes "sorted_on r (y # zs)" shows "sorted_on r (x # y # zs)" using assms \Linear_order r\ unfolding order_on_defs by (auto elim: transE intro: FieldI1) lemma sorted_on_Cons: shows "sorted_on r (x # xs) \ (x \ Field r \ sorted_on r xs \ (\y\set xs. (x, y) \ r))" using \Linear_order r\ unfolding order_on_defs by (induct xs arbitrary: x) (auto elim: transE) lemma sorted_on_distinct_set_unique: assumes "sorted_on r xs" "distinct xs" "sorted_on r ys" "distinct ys" "set xs = set ys" shows "xs = ys" proof - from assms have 1: "length xs = length ys" by (auto dest!: distinct_card) from assms show ?thesis proof(induct rule: list_induct2[OF 1]) case (2 x xs y ys) with \Linear_order r\ show ?case unfolding order_on_defs by (simp add: sorted_on_Cons) (metis antisymD insertI1 insert_eq_iff) qed simp qed lemma set_insort_on: shows "set (insort_key_on r f x xs) = insert x (set xs)" by (induct xs) auto lemma sort_key_on_simps [simp]: shows "sort_key_on r f [] = []" "sort_key_on r f (x#xs) = insort_key_on r f x (sort_key_on r f xs)" by (simp_all add: sort_key_on_def) lemma set_sort_on [simp]: shows "set (sort_key_on r f xs) = set xs" by (induct xs) (simp_all add: set_insort_on) lemma distinct_insort_on: shows "distinct (insort_key_on r f x xs) = (x \ set xs \ distinct xs)" by(induct xs) (auto simp: set_insort_on) lemma distinct_sort_on [simp]: shows "distinct (sort_key_on r f xs) = distinct xs" by (induct xs) (simp_all add: distinct_insort_on) lemma sorted_on_insort_key_on: assumes "f ` set (x # xs) \ Field r" shows "sorted_on r (map f (insort_key_on r f x xs)) = sorted_on r (map f xs)" using assms proof(induct xs) case (Cons x xs) with \Linear_order r\ show ?case unfolding order_on_defs by (auto 4 4 simp: sorted_on_Cons sorted_on_many set_insort_on refl_on_def total_on_def elim: transE) qed simp lemma sorted_on_insort_on: assumes "set (x # xs) \ Field r" shows "sorted_on r (insort_on r x xs) = sorted_on r xs" using sorted_on_insort_key_on[where f="\x. x"] assms by simp theorem sorted_on_sort_key_on [simp]: assumes "f ` set xs \ Field r" shows "sorted_on r (map f (sort_key_on r f xs))" using assms by (induct xs) (simp_all add: sorted_on_insort_key_on) theorem sorted_on_sort_on [simp]: assumes "set xs \ Field r" shows "sorted_on r (sort_on r xs)" using sorted_on_sort_key_on[where f="\x. x"] assms by simp lemma finite_sorted_on_distinct_unique: assumes "A \ Field r" assumes "finite A" shows "\!xs. set xs = A \ sorted_on r xs \ distinct xs" proof - from \finite A\ obtain xs where "set xs = A \ distinct xs" using finite_distinct_list by blast with \A \ Field r\ show ?thesis by (fastforce intro!: ex1I[where a="sort_on r xs"] simp: sorted_on_distinct_set_unique) qed end lemma sorted_on_linord_of_list_subseteq_r: assumes "Linear_order r" assumes "sorted_on r xs" assumes "distinct xs" shows "linord_of_list (rev xs) \ r" using assms proof(induct xs) case (Cons x xs) then have "linord_of_list (rev xs) \ r" by (simp add: sorted_on_Cons) with Cons.prems show ?case by (clarsimp simp: linord_of_list_append linord_of_list_singleton sorted_on_Cons) (meson contra_subsetD subsetI underS_incl_iff) qed simp lemma sorted_on_linord_of_list: assumes "Linear_order r" assumes "set xs = Field r" assumes "sorted_on r xs" assumes "distinct xs" shows "linord_of_list (rev xs) = r" proof(rule equalityI) from assms show "linord_of_list (rev xs) \ r" using sorted_on_linord_of_list_subseteq_r by blast next { fix x y assume xy: "(x, y) \ r" with \Linear_order r\ have "(y, x) \ r - Id" using Linear_order_in_diff_Id by (fastforce intro: FieldI1) with linord_of_list_Linear_order[of "rev xs" "Field r"] assms xy have "(x, y) \ linord_of_list (rev xs)" by simp (metis Diff_subset FieldI1 FieldI2 Linear_order_in_diff_Id linord_of_list_Field set_rev sorted_on_linord_of_list_subseteq_r subset_eq) } then show "r \ linord_of_list (rev xs)" by clarsimp qed lemma linord_of_listP_rev: assumes "z # zs \ set (subseqs xs)" assumes "y \ set zs" shows "linord_of_listP z y (rev xs)" using assms by (induct xs) (auto simp: Let_def linord_of_listP_append dest: subseqs_set) lemma linord_of_list_sorted_on_subseqs: assumes "ys \ set (subseqs xs)" assumes "distinct xs" shows "sorted_on (linord_of_list (rev xs)) ys" using assms proof(induct ys) case (Cons y ys) then show ?case using linord_of_list_Linear_order[where xs="rev xs" and ys="Field (linord_of_list (rev xs))"] by (force simp: Cons_in_subseqsD sorted_on_Cons linord_of_list_linord_of_listP linord_of_listP_rev dest: subseqs_set) qed simp lemma linord_of_list_sorted_on: assumes "distinct xs" shows "sorted_on (linord_of_list (rev xs)) xs" by (rule linord_of_list_sorted_on_subseqs[OF subseqs_refl \distinct xs\]) (*>*) lemma linear_order_on_list: assumes "linear_order_on ys r" assumes "ys = Field r" assumes "finite ys" shows "\!xs. r = linord_of_list xs \ distinct xs \ set xs = ys" using %invisible finite_sorted_on_distinct_unique[of r ys] sorted_on_linord_of_list[of r] assms by simp (metis distinct_rev linord_of_list_sorted_on rev_rev_ident set_rev) (*<*) end (*>*) diff --git a/thys/Stable_Matching/Contracts.thy b/thys/Stable_Matching/Contracts.thy --- a/thys/Stable_Matching/Contracts.thy +++ b/thys/Stable_Matching/Contracts.thy @@ -1,2677 +1,2680 @@ (*<*) theory Contracts imports Choice_Functions "HOL-Library.Dual_Ordered_Lattice" "HOL-Library.Bourbaki_Witt_Fixpoint" "HOL-Library.While_Combinator" "HOL-Library.Product_Order" begin (*>*) section\ \citet{HatfieldMilgrom:2005}: Matching with contracts \label{sec:contracts} \ text\ We take the original paper on matching with contracts by \citet{HatfieldMilgrom:2005} as our roadmap, which follows a similar path to \citet[\S2.5]{RothSotomayor:1990}. We defer further motivation to these texts. Our first move is to capture the scenarios described in their {\S}I(A) (p916) in a locale. \ locale Contracts = fixes Xd :: "'x::finite \ 'd::finite" fixes Xh :: "'x \ 'h::finite" fixes Pd :: "'d \ 'x rel" fixes Ch :: "'h \ 'x cfun" assumes Pd_linear: "\d. Linear_order (Pd d)" assumes Pd_range: "\d. Field (Pd d) \ {x. Xd x = d}" assumes Ch_range: "\h. \X. Ch h X \ {x\X. Xh x = h}" assumes Ch_singular: "\h. \X. inj_on Xd (Ch h X)" begin text \ The set of contracts is modelled by the type @{typ "'x"}, a free type variable that will later be interpreted by some non-empty set. Similarly @{typ "'d"} and @{typ "'h"} track the names of doctors and hospitals respectively. All of these are finite by virtue of belonging to the \finite\ type class. We fix four constants: \begin{itemize} \item \Xd\ (\Xh\) projects the name of the relevant doctor (hospital) from a contract; \item \Pd\ maps doctors to their linear preferences over some subset of contracts that name them (assumptions @{thm [source] Pd_linear} and @{thm [source] Pd_range}); and \item \Ch\ maps hospitals to their choice functions (\S\ref{sec:cf}), which are similarly constrained to contracts that name them (assumption @{thm [source] Ch_range}). Moreover their choices must name each doctor at most once, i.e., \Xd\ must be injective on these (assumption @{thm [source] "Ch_singular"}). \end{itemize} The reader familiar with the literature will note that we do not have a null contract (also said to represent the @{emph \outside option\} of unemployment), and instead use partiality of the doctors' preferences. This provides two benefits: firstly, \Xh\ is a total function here, and secondly we achieve some economy of description when instantiating this locale as \Pd\ only has to rank the relevant contracts. We note in passing that neither the doctors' nor hospitals' choice functions are required to be decisive, unlike the standard literature on choice functions (\S\ref{sec:cf}). In addition to the above, the following constitute the definitions that must be trusted for the results we prove to be meaningful. \ definition Cd :: "'d \ 'x cfun" where "Cd d \ set_option \ MaxR.MaxR_opt (Pd d)" definition CD_on :: "'d set \ 'x cfun" where "CD_on ds X = (\d\ds. Cd d X)" abbreviation CD :: "'x set \ 'x set" where "CD \ CD_on UNIV" definition CH :: "'x cfun" where "CH X = (\h. Ch h X)" text\ The function @{const "Cd"} constructs a choice function from the doctor's linear preferences (see \S\ref{sec:cf-linear}). Both @{const "CD"} and @{const "CH"} simply aggregate opinions in the obvious way. The functions @{const "CD_on"} is parameterized with a set of doctors to support the proofs of \S\ref{sec:contracts-vacancy-chain}. We also define \RD\ (\Rh\, \RH\) to be the subsets of a given set of contracts that are rejected by the doctors (hospitals). (The abbreviation @{const "Rf"} is defined in \S\ref{sec:cf-rf}.) \ abbreviation (input) RD_on :: "'d set \ 'x cfun" where "RD_on ds \ Rf (CD_on ds)" abbreviation (input) RD :: "'x cfun" where "RD \ RD_on UNIV" abbreviation (input) Rh :: "'h \ 'x cfun" where "Rh h \ Rf (Ch h)" abbreviation (input) RH :: "'x cfun" where "RH \ Rf CH" text \ A @{emph \mechanism\} maps doctor and hospital preferences into a match (here a set of contracts). \ type_synonym (in -) ('d, 'h, 'x) mechanism = "('d \ 'x rel) \ ('h \ 'x cfun) \ 'd set \ 'x set" (*<*) (* Pd *) lemmas Pd_linear' = Pd_linear[rule_format] lemmas Pd_range' = subsetD[OF Pd_range[rule_format], simplified, of x d] for x d lemma Pd_refl: assumes "x \ Field (Pd d)" shows "(x, x) \ Pd d" using assms Pd_linear' by (meson subset_refl underS_incl_iff) lemma Pd_Xd: assumes "(x, y) \ Pd d" shows "Xd x = d \ Xd y = d" using assms Pd_range contra_subsetD unfolding Field_def by blast lemma Above_Pd_Xd: assumes "x \ Above (Pd d) X" shows "Xd x = d" using assms by (blast dest: Above_Field Pd_range') lemma AboveS_Pd_Xd: assumes "x \ AboveS (Pd d) X" shows "Xd x = d" using assms by (blast dest: AboveS_Field Pd_range') (* Cd *) interpretation Cd: linear_cf "Pd d" "Cd d" for d using Cd_def Pd_linear by unfold_locales simp_all lemmas Cd_domain = Cd.domain lemmas Cd_f_range = Cd.f_range lemmas Cd_range = Cd.range lemmas Cd_range' = Cd.range' lemmas Rf_Cd_mono = Cd.Rf_mono_on[of UNIV, unfolded mono_on_mono] lemmas Cd_Chernoff = Cd.Chernoff lemmas Cd_path_independent = Cd.path_independent lemmas Cd_iia = Cd.iia lemmas Cd_irc = Cd.irc lemmas Cd_lad = Cd.lad lemmas Cd_mono = Cd.mono lemmas Cd_greatest = Cd.greatest lemmas Cd_preferred = Cd.preferred lemmas Cd_singleton = Cd.singleton lemmas Cd_union = Cd.union lemmas Cd_idem = iia_f_idem[OF Cd.f_range[of UNIV d, folded Cd_def] Cd_iia[of UNIV], simplified] for d lemma Cd_Xd: shows "x \ Cd d X \ Xd x = d" using Pd_range Cd_range by fastforce lemma Cd_inj_on_Xd: shows "inj_on Xd (Cd d X)" by (rule inj_onI) (clarsimp simp: Cd_Xd Cd_singleton) lemma Cd_range_disjoint: assumes "d \ d'" shows "Cd d A \ Cd d' A = {}" using assms Cd_range Pd_range by blast lemma Cd_single: assumes "x \ X" assumes "inj_on Xd X" assumes "x \ Field (Pd d)" shows "x \ Cd d X" using assms Pd_linear unfolding Cd_greatest greatest_def by clarsimp (metis Pd_Xd inj_on_eq_iff subset_refl underS_incl_iff) lemma Cd_Above: shows "Cd d X = Above (Pd d) (X \ Field (Pd d)) \ X" unfolding Cd_greatest greatest_Above Above_def by blast (* Code generator setup. Repeats a lot of stuff. *) definition maxR :: "'d \ 'x \ 'x \ 'x" where "maxR d x y = (if (x, y) \ Pd d then y else x)" definition MaxR_f :: "'d \ 'x \ 'x option \ 'x option" where "MaxR_f d = (\x acc. if x \ Field (Pd d) then Some (case acc of None \ x | Some y \ maxR d x y) else acc)" lemma MaxR_maxR: shows "MaxR.maxR (Pd d) = maxR d" by (simp add: fun_eq_iff maxR_def Cd.maxR_code) lemma MaxR_MaxR_f: shows "MaxR.MaxR_f (Pd d) = MaxR_f d" by (simp add: fun_eq_iff Cd.MaxR_f_code MaxR_f_def MaxR_maxR cong: option.case_cong) lemmas Cd_code[code] = Cd.code[unfolded MaxR_MaxR_f] lemma Cd_simps[simp, nitpick_simp]: shows "Cd d {} = {}" "Cd d (insert x A) = (if x \ Field (Pd d) then if Cd d A = {} then {x} else {maxR d x y |y. y \ Cd d A} else Cd d A)" unfolding Cd.simps MaxR_maxR by simp_all (* CD *) lemma CD_on_def2: shows "CD_on ds A = (\d\ds. Cd d (A \ Field (Pd d)))" using Cd_domain unfolding CD_on_def by blast lemma CD_on_Xd: assumes "x \ CD_on ds A" shows "Xd x \ ds" using assms Cd_Xd unfolding CD_on_def by blast lemma mem_CD_on_Cd: shows "x \ CD_on ds X \ (x \ Cd (Xd x) X \ Xd x \ ds)" unfolding CD_on_def using Cd_range Cd_Xd by blast lemma CD_on_domain: assumes "d \ ds" shows "CD_on ds A \ Field (Pd d) = Cd d (A \ Field (Pd d))" unfolding CD_on_def2 using assms Cd_range by (force dest: Pd_range') lemma CD_on_range: shows "CD_on ds A \ A \ (\d\ds. Field (Pd d))" using Cd_range unfolding CD_on_def by blast lemmas CD_on_range' = subsetD[OF CD_on_range] lemma CD_on_f_range_on: shows "f_range_on A (CD_on ds)" by (rule f_range_onI) (meson CD_on_range Int_subset_iff) lemma RD_on_mono: shows "mono (RD_on ds)" unfolding CD_on_def by (rule monoI) (auto dest: monoD[OF Rf_Cd_mono]) lemma CD_on_Chernoff: shows "Chernoff (CD_on ds)" using mono_on_mono RD_on_mono[of ds] Rf_mono_on_iia_on[of UNIV] Chernoff_on_iia_on by (simp add: fun_eq_iff) blast lemma CD_on_irc: shows "irc (CD_on ds)" by (rule ircI) (fastforce simp: CD_on_def ircD[OF Cd_irc] simp del: Cd_simps cong: SUP_cong) lemmas CD_on_consistency = irc_on_consistency_on[OF CD_on_irc, simplified] lemma CD_on_path_independent: shows "path_independent (\X. CD_on ds X)" using CD_on_f_range_on CD_on_Chernoff CD_on_consistency by (blast intro: path_independent_onI2) lemma CD_on_simps: shows "CD_on ds {} = {}" using CD_on_range by blast lemmas CD_on_iia = RD_on_mono[unfolded Rf_mono_iia] lemmas CD_on_idem = iia_f_idem[OF CD_on_f_range_on CD_on_iia, simplified] lemma CD_on_inj_on_Xd: shows "inj_on Xd (CD_on ds X)" unfolding CD_on_def by (rule inj_onI) (clarsimp simp: Cd_Xd Cd_singleton) lemma CD_on_card: shows "card (CD_on ds X) = (\d\ds. card (Cd d X))" unfolding CD_on_def by (simp add: card_UN_disjoint Cd_range_disjoint) lemma CD_on_closed: assumes "inj_on Xd X" assumes "X \ (\d\ds. Field (Pd d))" shows "CD_on ds X = X" using assms Cd_domain Cd_single[OF _ assms(1)] unfolding CD_on_def2 by (force dest: Cd_range') (* Ch *) lemmas Ch_singular' = Ch_singular[rule_format] lemmas Ch_range' = subsetD[OF Ch_range[rule_format], simplified, of x h X] for x h X lemma Ch_simps: shows "Ch h {} = {}" using Ch_range by blast lemma Ch_range_disjoint: assumes "h \ h'" shows "Ch h A \ Ch h' A = {}" using assms Ch_range by blast lemma Ch_f_range: shows "f_range (Ch h)" using Ch_range unfolding f_range_on_def by blast (* CH *) lemma CH_card: shows "card (CH X) = (\h\UNIV. card (Ch h X))" unfolding CH_def by (simp add: card_UN_disjoint Ch_range_disjoint) lemma CH_simps: shows "CH {} = {}" unfolding CH_def by (simp add: Ch_simps) lemma CH_range: shows "CH A \ A" unfolding CH_def using Ch_range by blast lemmas CH_range' = subsetD[OF CH_range] lemmas CH_f_range_on = f_range_onI[OF CH_range] lemma mem_CH_Ch: shows "x \ CH X \ x \ Ch (Xh x) X" unfolding CH_def using Ch_range by blast lemma mem_Ch_CH: assumes "x \ Ch h X" shows "x \ CH X" unfolding CH_def using assms Ch_range by blast (*>*) text\ An @{emph \allocation\} is a set of contracts where each names a distinct doctor. (Hospitals can contract multiple doctors.) \ abbreviation (input) allocation :: "'x set \ bool" where "allocation \ inj_on Xd" text\ We often wish to extract a doctor's or a hospital's contract from an @{const "allocation"}. \ definition dX :: "'x set \ 'd \ 'x set" where "dX X d = {x \ X. Xd x = d}" definition hX :: "'x set \ 'h \ 'x set" where "hX X h = {x \ X. Xh x = h}" (*<*) lemma dX_union: shows "dX (X \ Y) d = dX X d \ dX Y d" unfolding dX_def by auto lemma dX_range: shows "\d. dX X d \ {x. Xd x = d}" unfolding dX_def by clarsimp lemma dX_range': assumes "x \ dX X d" shows "x \ X \ Xd x = d" using assms unfolding dX_def by simp lemma dX_empty_or_singleton: assumes "allocation X" shows "\d. dX X d = {} \ (\x. dX X d = {x})" unfolding dX_def using \allocation X\ by (fastforce dest: inj_onD) lemma dX_linear: assumes "allocation X" shows "Linear_order (dX X d \ dX X d)" using spec[OF dX_empty_or_singleton[OF \allocation X\], where x=d] by fastforce lemma dX_singular: assumes "allocation X" assumes "x \ X" assumes "d = Xd x" shows "dX X d = {x}" using assms unfolding dX_def by (fastforce dest: inj_onD) lemma dX_Int_Field_Pd: assumes "dX X d \ Field (Pd d)" shows "X \ Field (Pd d) = dX X d" using assms unfolding dX_def by (fastforce dest: Pd_range') lemma Cd_Above_dX: assumes "dX X d \ Field (Pd d)" shows "Cd d X = Above (Pd d) (dX X d) \ X" using assms unfolding Cd_greatest greatest_Above Above_def dX_def by (auto dest: Pd_range') (*>*) text\ @{emph \Stability\} is the key property we look for in a match (here a set of contracts), and consists of two parts. Firstly, we ask that it be @{emph \individually rational\}, i.e., the contracts in the match are actually acceptable to all participants. Note that this implies the match is an @{const "allocation"}. \ definition individually_rational_on :: "'d set \ 'x set \ bool" where "individually_rational_on ds X \ CD_on ds X = X \ CH X = X" abbreviation individually_rational :: "'x set \ bool" where "individually_rational \ individually_rational_on UNIV" text\ The second condition requires that there be no coalition of a hospital and one or more doctors who prefer another set of contracts involving them; the hospital strictly, the doctors weakly. Contrast this definition with the classical one for stable marriages given in \S\ref{sec:sotomayor}. \ definition blocking_on :: "'d set \ 'x set \ 'h \ 'x set \ bool" where "blocking_on ds X h X' \ X' \ Ch h X \ X' = Ch h (X \ X') \ X' \ CD_on ds (X \ X')" definition stable_no_blocking_on :: "'d set \ 'x set \ bool" where "stable_no_blocking_on ds X \ (\h X'. \blocking_on ds X h X')" abbreviation stable_no_blocking :: "'x set \ bool" where "stable_no_blocking \ stable_no_blocking_on UNIV" definition stable_on :: "'d set \ 'x set \ bool" where "stable_on ds X \ individually_rational_on ds X \ stable_no_blocking_on ds X" abbreviation stable :: "'x set \ bool" where "stable \ stable_on UNIV" (*<*) lemma stable_onI: assumes "individually_rational_on ds X" assumes "stable_no_blocking_on ds X" shows "stable_on ds X" unfolding stable_on_def using assms by blast lemma individually_rational_onI: assumes "CD_on ds X = X" assumes "CH X = X" shows "individually_rational_on ds X" unfolding individually_rational_on_def using assms by blast lemma individually_rational_on_CD_on: assumes "individually_rational_on ds X" shows "CD_on ds X = X" using assms unfolding individually_rational_on_def by blast lemma individually_rational_on_Cd: assumes "individually_rational_on ds X" shows "Cd d X = dX X d" using individually_rational_on_CD_on[OF assms] by (auto simp: dX_def mem_CD_on_Cd dest: Cd_range' Cd_Xd) lemma individually_rational_on_empty: shows "individually_rational_on ds {}" by (simp add: CD_on_simps CH_simps individually_rational_onI) lemma blocking_onI: assumes "X'' \ Ch h X" assumes "X'' = Ch h (X \ X'')" assumes "\x. x \ X'' \ x \ CD_on ds (X \ X'')" shows "blocking_on ds X h X''" unfolding blocking_on_def using assms by blast lemma blocking_on_imp_not_stable: assumes "blocking_on ds X h X''" shows "\stable_on ds X" unfolding stable_on_def stable_no_blocking_on_def using assms by blast lemma blocking_on_allocation: assumes "blocking_on ds X h X''" shows "allocation X''" using assms unfolding blocking_on_def by (metis Ch_singular') lemma blocking_on_Field: assumes "blocking_on ds X h X''" shows "dX X'' d \ Field (Pd d)" using assms blocking_on_allocation[OF assms] unfolding blocking_on_def dX_def by (force simp: Pd_range' dest: CD_on_range') lemma blocking_on_CD_on: assumes "blocking_on ds X h X''" shows "X'' \ CD_on ds (X \ X'')" using assms unfolding blocking_on_def by blast lemma blocking_on_CD_on': assumes "blocking_on ds X h X''" assumes "x \ X''" shows "x \ CD_on ds (X \ X'')" using assms unfolding blocking_on_def by blast lemma blocking_on_Cd: assumes "blocking_on ds X h X''" shows "dX X'' d \ Cd d (X \ X'')" using assms unfolding blocking_on_def by (force dest: dX_range' simp: mem_CD_on_Cd) lemma stable_no_blocking_onI: assumes "\h X''. \X'' = Ch h (X \ X''); X'' \ Ch h X; X'' \ CD_on ds (X \ X'')\ \ False" shows "stable_no_blocking_on ds X" unfolding stable_no_blocking_on_def blocking_on_def using assms by blast lemma stable_no_blocking_onI2: assumes "\h X''. blocking_on ds X h X'' \ False" shows "stable_no_blocking_on ds X" unfolding stable_no_blocking_on_def using assms by blast lemma "stable_no_blocking_on ds UNIV" using stable_no_blocking_onI by fastforce lemma assumes "stable_on ds X" shows stable_on_CD_on: "CD_on ds X = X" and stable_on_Xd: "x \ X \ Xd x \ ds" and stable_on_range': "x \ X \ x \ Field (Pd (Xd x))" and stable_on_CH: "CH X = X" and stable_on_no_blocking_on: "stable_no_blocking_on ds X" using assms mem_CD_on_Cd Cd_range' Pd_range' unfolding stable_on_def individually_rational_on_def by blast+ lemma stable_on_allocation: assumes "stable_on ds X" shows "allocation X" using assms unfolding stable_on_def individually_rational_on_def by (metis CD_on_inj_on_Xd) lemma stable_on_blocking_onD: assumes "stable_on ds X" shows "\X'' = Ch h (X \ X''); X'' \ CD_on ds (X \ X'')\ \ X'' = Ch h X" using \stable_on ds X\ unfolding stable_on_def individually_rational_on_def stable_no_blocking_on_def blocking_on_def by blast lemma not_stable_on_cases[consumes 1, case_names not_individually_rational not_no_blocking]: assumes "\ stable_on ds X" assumes "\ individually_rational_on ds X \ P" assumes "\ stable_no_blocking_on ds X \ P" shows "P" using assms unfolding stable_on_def by blast (*>*) text\\ end subsection\ Theorem~1: Existence of stable pairs \ text\ We proceed to define a function whose fixed points capture all stable matches. \citet[I(B), p917]{HatfieldMilgrom:2005} provide the following intuition: \begin{quote} The first theorem states that a set of contracts is stable if any alternative contract would be rejected by some doctor or some hospital from its suitably defined opportunity set. In the formulas below, think of the doctors' opportunity set as @{term "XD"} and the hospitals' opportunity set as @{term "XH"}. If @{term "X'"} is the corresponding stable set, then @{term "XD"} must include, in addition to @{term "X'"}, all contracts that would not be rejected by the hospitals, and @{term "XH"} must similarly include @{term "X'"} and all contracts that would not be rejected by the doctors. If @{term "X'"} is stable, then every alternative contract is rejected by somebody, so @{term "X = XH \ XD"} [where @{term "X"} is the set of all contracts]. This logic is summarized in the first theorem. \end{quote} See also \citet[p6,\S4]{Fleiner:2003} and \citet[\S2]{Fleiner:2002}, from whom we adopt the term @{emph \stable pair\}. \ context Contracts begin definition stable_pair_on :: "'d set \ 'x set \ 'x set \ bool" where "stable_pair_on ds = (\(XD, XH). XD = - RH XH \ XH = - RD_on ds XD)" abbreviation stable_pair :: "'x set \ 'x set \ bool" where "stable_pair \ stable_pair_on UNIV" abbreviation match :: "'x set \ 'x set \ 'x set" where "match X \ fst X \ snd X" text \ \citet[Theorem~1]{HatfieldMilgrom:2005} state that every solution @{term "(XD, XH)"} of @{const "stable_pair"} yields a stable match @{term "XD \ XH"}, and conversely, i.e., every stable match is the intersection of some stable pair. \citet{AygunSonmez:2012-WP2} show that neither is the case without further restrictions on the hospitals' choice functions @{term "Ch"}; we exhibit their counterexample below. Even so we can establish some properties in the present setting: \ lemma stable_pair_on_CD_on: assumes "stable_pair_on ds XD_XH" shows "match XD_XH = CD_on ds (fst XD_XH)" using %invisible assms CD_on_range unfolding stable_pair_on_def split_def fst_conv snd_conv by blast lemma stable_pair_on_CH: assumes "stable_pair_on ds XD_XH" shows "match XD_XH = CH (snd XD_XH)" using %invisible assms CH_range unfolding stable_pair_on_def split_def fst_conv snd_conv by blast lemma stable_pair_on_CD_on_CH: assumes "stable_pair_on ds XD_XH" shows "CD_on ds (fst XD_XH) = CH (snd XD_XH)" using %invisible assms stable_pair_on_CD_on stable_pair_on_CH by blast lemma stable_pair_on_allocation: assumes "stable_pair_on ds XD_XH" shows "allocation (match XD_XH)" unfolding %invisible stable_pair_on_CD_on[OF assms] by (rule CD_on_inj_on_Xd) (*<*) lemma stable_pair_onI: assumes "fst XD_XH = - RH (snd XD_XH)" assumes "snd XD_XH = - RD_on ds (fst XD_XH)" shows "stable_pair_on ds XD_XH" using assms unfolding stable_pair_on_def split_def by blast lemma stable_pair_onE: shows "\stable_pair_on ds XD_XH; \- RH (snd XD_XH) = fst XD_XH; - RD_on ds (fst XD_XH) = snd XD_XH\ \ P\ \ P" unfolding stable_pair_on_def split_def by blast lemma stable_pair_on_Cd: assumes "stable_pair_on ds XD_XH" assumes "d \ ds" shows "Cd d (fst XD_XH) = match XD_XH \ Field (Pd d)" using stable_pair_on_CD_on[OF \stable_pair_on ds XD_XH\] CD_on_domain Cd_domain \d \ ds\ by simp lemma stable_pair_on_Cd_match: assumes "stable_pair_on ds XD_XH" assumes "d \ ds" shows "Cd d (match XD_XH) = Cd d (fst XD_XH)" using assms by (metis Cd_domain Cd_idem stable_pair_on_Cd) lemma stable_pair_on_Xd: assumes "stable_pair_on ds XD_XH" assumes "x \ match XD_XH" shows "Xd x \ ds" using assms CD_on_Xd unfolding stable_pair_on_def split_def by blast lemma stable_pair_on_match_Cd: assumes "stable_pair_on ds XD_XH" assumes "x \ match XD_XH" shows "x \ Cd (Xd x) (match XD_XH)" using assms by (metis (full_types) CD_on_def Cd_Xd UN_iff stable_pair_on_CD_on stable_pair_on_Cd_match) (*>*) text\ We run out of steam on the following two lemmas, which are the remaining requirements for stability. \ lemma assumes "stable_pair_on ds XD_XH" shows "individually_rational_on ds (match XD_XH)" oops lemma assumes "stable_pair_on ds XD_XH" shows "stable_no_blocking (match XD_XH)" oops text\ \citet{HatfieldMilgrom:2005} also claim that the converse holds: \ lemma assumes "stable_on ds X" obtains XD_XH where "stable_pair_on ds XD_XH" and "X = match XD_XH" oops text\ Again, the following counterexample shows that the @{const substitutes} condition on @{term "Ch"} is too weak to guarantee this. We show it holds under stronger assumptions in \S\ref{sec:contracts-t1-converse}. \ end subsubsection\ Theorem~1 does not hold \citep{AygunSonmez:2012-WP2} \label{sec:contracts-t1-counterexample} \ text\ The following counterexample, due to \citet[\S3: Example~2]{AygunSonmez:2012-WP2}, comprehensively demonstrates that \citet[Theorem~1]{HatfieldMilgrom:2005} does not hold. We create three types: \D2\ consists of two elements, representing the doctors, and \H\ is the type of the single hospital. There are four contracts in the type \X4\. \ datatype D2 = D1 | D2 datatype H1 = H datatype X4 = Xd1 | Xd1' | Xd2 | Xd2' (*<*) lemma D2_UNIV: shows "UNIV = set [D1, D2]" using D2.exhaust by auto instantiation D2 :: enum begin definition "enum_class.enum = [D1, D2]" definition "enum_class.enum_all P = (P D1 \ P D2)" definition "enum_class.enum_ex P = (P D1 \ P D2)" instance by standard (simp_all add: enum_D2_def enum_all_D2_def enum_ex_D2_def D2_UNIV) end lemma D2_ALL: shows "(\d. P d) = (\d\{D1, D2}. P d)" using D2_UNIV by auto lemma D2_UNION: shows "(\d. P d) = (\d\{D1, D2}. P d)" using D2_UNIV by auto instance H1 :: finite by standard (metis (full_types) H1.exhaust ex_new_if_finite finite.intros(1) finite_insert insert_subset subset_insertI) lemma X4_UNIV: shows "UNIV = set [Xd1, Xd1', Xd2, Xd2']" using X4.exhaust by auto lemmas X4_pow = subset_subseqs[OF subset_trans[OF subset_UNIV Set.equalityD1[OF X4_UNIV]]] instance X4 :: finite by standard (simp add: X4_UNIV) lemma X4_ALL: shows "(\X''. P X'') \ (\X''\set ` set (subseqs [Xd1, Xd1', Xd2, Xd2']). P X'')" using X4_pow by blast (*>*) primrec X4d :: "X4 \ D2" where "X4d Xd1 = D1" | "X4d Xd1' = D1" | "X4d Xd2 = D2" | "X4d Xd2' = D2" abbreviation X4h :: "X4 \ H1" where "X4h _ \ H" primrec PX4d :: "D2 \ X4 rel" where "PX4d D1 = linord_of_list [Xd1', Xd1]" | "PX4d D2 = linord_of_list [Xd2, Xd2']" function CX4h :: "H1 \ X4 cfun" where "CX4h _ {Xd1} = {Xd1}" | "CX4h _ {Xd1'} = {Xd1'}" | "CX4h _ {Xd2} = {Xd2}" | "CX4h _ {Xd2'} = {Xd2'}" | "CX4h _ {Xd1, Xd1'} = {Xd1}" | "CX4h _ {Xd1, Xd2} = {Xd1, Xd2}" | "CX4h _ {Xd1, Xd2'} = {Xd2'}" | "CX4h _ {Xd1', Xd2} = {Xd1'}" | "CX4h _ {Xd1', Xd2'} = {Xd1', Xd2'}" | "CX4h _ {Xd2, Xd2'} = {Xd2}" | "CX4h _ {Xd1, Xd1', Xd2} = {}" | "CX4h _ {Xd1, Xd1', Xd2'} = {}" | "CX4h _ {Xd1, Xd2, Xd2'} = {}" | "CX4h _ {Xd1', Xd2, Xd2'} = {}" | "CX4h _ {Xd1, Xd1', Xd2, Xd2'} = {}" | "CX4h _ {} = {}" apply %invisible (case_tac x) apply (cut_tac X=b in X4_pow) apply auto done (*<*) termination by %invisible lexicographic_order lemma PX4d_linear: shows "Linear_order (PX4d d)" by (cases d) (simp_all add: linord_of_list_Linear_order) lemma PX4d_range: shows "Field (PX4d d) \ {x. X4d x = d}" by (cases d) simp_all lemma CX4h_range: shows "CX4h h X \ {x \ X. H = h}" by (cases "(h, X)" rule: CX4h.cases) (auto simp: spec[OF H1.nchotomy, of h]) lemma CX4h_singular: shows "inj_on X4d (CX4h h X)" by (cases "(h, X)" rule: CX4h.cases) auto (*>*) text\\ interpretation StableNoDecomp: Contracts X4d X4h PX4d CX4h using %invisible PX4d_linear PX4d_range CX4h_range CX4h_singular by unfold_locales blast+ text\ There are two stable matches in this model. \ (*<*) lemma Xd1_Xd2_stable: shows "StableNoDecomp.stable {Xd1, Xd2}" proof(rule StableNoDecomp.stable_onI) show "StableNoDecomp.individually_rational {Xd1, Xd2}" by (simp add: StableNoDecomp.individually_rational_on_def StableNoDecomp.CD_on_def StableNoDecomp.CH_def insert_commute D2_UNION cong add: SUP_cong_simp) show "StableNoDecomp.stable_no_blocking {Xd1, Xd2}" apply (rule StableNoDecomp.stable_no_blocking_onI) apply (rule_tac x="(H, X'')" in CX4h.cases) apply (simp_all add: insert_commute) done qed lemma Xd1'_Xd2'_stable: shows "StableNoDecomp.stable {Xd1', Xd2'}" proof(rule StableNoDecomp.stable_onI) show "StableNoDecomp.individually_rational {Xd1', Xd2'}" by (simp add: StableNoDecomp.individually_rational_on_def StableNoDecomp.CD_on_def StableNoDecomp.CH_def insert_commute D2_UNION cong add: SUP_cong_simp) show "StableNoDecomp.stable_no_blocking {Xd1', Xd2'}" apply (rule StableNoDecomp.stable_no_blocking_onI) apply (rule_tac x="(H, X'')" in CX4h.cases) apply (simp_all add: insert_commute) done qed (*>*) text\\ lemma stable: shows "StableNoDecomp.stable X \ X = {Xd1, Xd2} \ X = {Xd1', Xd2'}" (*<*) (is "?lhs = ?rhs") proof(rule iffI) assume ?lhs then show ?rhs using X4_pow[where X=X] unfolding StableNoDecomp.stable_on_def StableNoDecomp.individually_rational_on_def StableNoDecomp.stable_no_blocking_on_def StableNoDecomp.blocking_on_def StableNoDecomp.CD_on_def StableNoDecomp.CH_def by simp (elim disjE, simp_all add: D2_UNION X4_ALL insert_commute StableNoDecomp.maxR_def cong add: SUP_cong_simp) next assume ?rhs then show ?lhs using Xd1_Xd2_stable Xd1'_Xd2'_stable by blast qed (*>*) text\ However neither arises from a pair \XD, XH\ that satisfy @{const "StableNoDecomp.stable_pair"}: \ + + + lemma StableNoDecomp_XD_XH: shows "StableNoDecomp.stable_pair (XD, XH) \ (XD = {} \ XH = {Xd1, Xd1', Xd2, Xd2'})" (*<*) (is "?lhs = ?rhs") proof(rule iffI) note image_cong_simp [cong del] note INF_cong_simp [cong] note SUP_cong_simp [cong] assume ?lhs then show ?rhs (* Expand the Cartesian product and check. *) using X4_pow [of XD] X4_pow [of XH] apply simp apply (erule StableNoDecomp.stable_pair_onE) - apply (elim disjE) + apply (elim disjE) apply (simp_all add: StableNoDecomp.CD_on_def StableNoDecomp.CH_def) unfolding X4_UNIV [simplified] apply (auto simp: D2_ALL D2_UNION X4_ALL insert_commute StableNoDecomp.maxR_def linord_of_list_linord_of_listP) done next assume ?rhs then show ?lhs unfolding StableNoDecomp.stable_pair_on_def using X4.exhaust by (auto simp: StableNoDecomp.CH_def) qed (*>*) text\\ proposition assumes "StableNoDecomp.stable_pair (XD, XH)" shows "\StableNoDecomp.stable (XD \ XH)" using %invisible assms apply (subst (asm) StableNoDecomp_XD_XH) apply (simp add: StableNoDecomp.stable_on_def StableNoDecomp.stable_no_blocking_on_def StableNoDecomp.blocking_on_def StableNoDecomp.individually_rational_on_empty) apply (auto simp: StableNoDecomp.mem_CD_on_Cd MaxR_def exI[where x=D1] exI[where x=H] exI[where x="{Xd1}"]) done text\ Moreover the converse of Theorem~1 does not hold either: the single decomposition that satisfies @{const "StableNoDecomp.stable_pair"} (@{thm [source] "StableNoDecomp_XD_XH"}) does not yield a stable match: \ proposition assumes "StableNoDecomp.stable X" shows "\(\XD XH. StableNoDecomp.stable_pair (XD, XH) \ X = XD \ XH)" using %invisible assms StableNoDecomp_XD_XH stable by fastforce text\ So there is not hope for \citet[Theorem~1]{HatfieldMilgrom:2005} as it stands. Note that the counterexample satisfies the @{const "substitutes"} condition (see \S\ref{sec:cf-substitutes}): \ lemma shows "substitutes (CX4h H)" proof %invisible (rule substitutes_onI) fix A a b assume "b \ CX4h H (insert b A)" then show "b \ CX4h H (insert a (insert b A))" apply (case_tac [!] a) apply (case_tac [!] b) apply ( (rule CX4h.cases[of "(H, A)"], auto simp: insert_commute)[1] )+ done qed text\ Therefore while @{const "substitutes"} supports the monotonicity argument that underpins their deferred-acceptance algorithm (see \S\ref{sec:contracts-algorithmics}), it is not enough to rescue Theorem~1. One way forward is to constrain the hospitals' choice functions, which we discuss in the next section. \ subsubsection\ Theorem 1 holds with @{emph \independence of rejected contracts\} \label{sec:contracts-irc} \ text\ \citet{AygunSonmez:2012-WP2} propose to rectify this issue by requiring hospitals' choices to satisfy @{const "irc"} (\S\ref{sec:cf-irc}). Reassuringly their counterexample fails to satisfy it: \ lemma shows "\irc (CX4h H)" by %invisible (fastforce simp: insert_commute dest: irc_onD[where a="Xd2" and B="{Xd1, Xd1'}"]) text\ We adopt this hypothesis by extending the @{const "Contracts"} locale: \ locale ContractsWithIRC = Contracts + assumes Ch_irc: "\h. irc (Ch h)" begin text\ This property requires that \Ch\ behave, for example, as follows: \ lemma Ch_domain: shows "Ch h (A \ {x. Xh x = h}) = Ch h A" using %invisible irc_on_discard[OF spec[OF Ch_irc, of h], where B="A \ {x. Xh x = h}" and C="A - {x. Xh x = h}"] by (fastforce simp: Un_Diff_Int ac_simps dest: Ch_range') lemma %invisible CH_domain: shows "CH A \ {x. Xh x = h} = Ch h (A \ {x. Xh x = h})" unfolding CH_def using Ch_range by (auto simp: Ch_domain) lemma %invisible stable_pair_on_Ch: assumes "stable_pair_on ds XD_XH" shows "Ch h (snd XD_XH) = match XD_XH \ {x. Xh x = h}" using stable_pair_on_CH[OF assms] CH_domain Ch_domain by simp lemmas %invisible Ch_consistency = irc_on_consistency_on[OF spec[OF Ch_irc], simplified, of h] for h lemmas Ch_irc_idem = consistency_on_f_idem[OF Ch_f_range Ch_consistency, simplified] lemma CH_irc_idem: shows "CH (CH A) = CH A" unfolding %invisible CH_def by (metis CH_def CH_domain Ch_domain Ch_irc_idem) lemma Ch_CH_irc_idem: shows "Ch h (CH A) = Ch h A" using %invisible CH_domain CH_irc_idem Ch_domain by blast text\ This suffices to show the left-to-right direction of Theorem~1. \ lemma stable_pair_on_individually_rational: assumes "stable_pair_on ds XD_XH" shows "individually_rational_on ds (match XD_XH)" by %invisible (metis CD_on_idem CH_irc_idem stable_pair_on_CD_on stable_pair_on_CD_on_CH assms individually_rational_onI) lemma stable_pair_on_stable_no_blocking_on: assumes "stable_pair_on ds XD_XH" shows "stable_no_blocking_on ds (match XD_XH)" proof(rule stable_no_blocking_onI) fix h X'' assume C: "X'' = Ch h (match XD_XH \ X'')" assume NE: "X'' \ Ch h (match XD_XH)" assume CD: "X'' \ CD_on ds (match XD_XH \ X'')" have "X'' \ snd XD_XH" proof - from CD have "X'' \ CD_on ds (CD_on ds (fst XD_XH) \ X'')" by (simp only: stable_pair_on_CD_on[OF assms]) then have "X'' \ CD_on ds (fst XD_XH \ X'')" using CD_on_path_independent unfolding path_independent_def by (simp add: Un_commute) moreover have "fst XD_XH \ CD_on ds (fst XD_XH \ X'') \ CD_on ds (fst XD_XH)" using CD_on_Chernoff unfolding Chernoff_on_def by (simp add: inf_commute) ultimately show ?thesis using assms unfolding stable_pair_on_def split_def by blast qed then have "Ch h (snd XD_XH) = Ch h (Ch h (snd XD_XH) \ X'')" by (force intro!: consistencyD[OF Ch_consistency] dest: Ch_range') moreover from NE have "X'' \ Ch h (snd XD_XH)" using stable_pair_on_CH[OF assms] CH_domain[of _ h] Ch_domain[of h] by (metis Ch_irc_idem) ultimately have "X'' \ Ch h (match XD_XH \ X'')" using stable_pair_on_CH[OF assms] CH_domain[of _ h] Ch_domain[of h] by (metis (no_types, lifting) inf.right_idem inf_sup_distrib2) with C show False by blast qed theorem stable_pair_on_stable_on: assumes "stable_pair_on ds XD_XH" shows "stable_on ds (match XD_XH)" using %invisible assms stable_pair_on_allocation stable_pair_on_individually_rational stable_pair_on_stable_no_blocking_on by (blast intro: stable_onI) end subsubsection\ The converse of Theorem~1 \label{sec:contracts-t1-converse} \ text (in Contracts) \ The forward direction of Theorem~1 gives us a way of finding stable matches by computing fixed points of a function closely related to @{const "stable_pair"} (see \S\ref{sec:contracts-algorithmics}). The converse says that every stable match can be decomposed in this way, which implies that the stable matches form a lattice (see also \S\ref{sec:contracts-algorithmics}). The following proofs assume that the hospitals' choice functions satisfy @{const "substitutes"} and @{const "irc"}. \ context ContractsWithIRC begin context fixes ds :: "'b set" fixes X :: "'a set" begin text\ Following \citet[Proof of Theorem~1]{HatfieldMilgrom:2005}, we partition the set of all contracts into @{term "[X, XD_smallest - X, XH_largest - X]"} with careful definitions of the two sets @{term "XD_smallest"} and @{term "XH_largest"}. Specifically @{term "XH_largest"} contains all contracts ranked at least as good as those in @{term "X"} by the doctors, considering unemployment and unacceptable contracts. Similarly @{term "XD_smallest"} contains those ranked at least as poorly. \ definition XH_largest :: "'a set" where "XH_largest = {y. Xd y \ ds \ y \ Field (Pd (Xd y)) \ (\x \ dX X (Xd y). (x, y) \ Pd (Xd y))}" definition XD_smallest :: "'a set" where "XD_smallest = - (XH_largest - X)" context assumes "stable_on ds X" begin lemma Ch_XH_largest_Field: assumes "x \ Ch h XH_largest" shows "x \ Field (Pd (Xd x))" using assms unfolding XH_largest_def by (blast dest: Ch_range') lemma Ch_XH_largest_Xd: assumes "x \ Ch h XH_largest" shows "Xd x \ ds" using assms unfolding XH_largest_def by (blast dest: Ch_range') lemma X_subseteq_XH_largest: shows "X \ XH_largest" proof(rule subsetI) fix x assume "x \ X" then obtain d where "d \ ds" "x \ Cd d X" using stable_on_CD_on[OF \stable_on ds X\] unfolding CD_on_def by blast with \stable_on ds X\ show "x \ XH_largest" using Pd_linear' Pd_range' Cd_range subset_Image1_Image1_iff[of "Pd d"] stable_on_allocation[of ds X] unfolding XH_largest_def linear_order_on_def partial_order_on_def stable_on_def inj_on_def dX_def by simp blast qed lemma X_subseteq_XD_smallest: shows "X \ XD_smallest" unfolding XD_smallest_def by blast lemma X_XD_smallest_XH_largest: shows "X = XD_smallest \ XH_largest" using X_subseteq_XH_largest unfolding XD_smallest_def by blast text\ The goal of the next few lemmas is to show the constituents of @{term "stable_pair_on ds (XD_smallest, XH_largest)"}. Intuitively, if a doctor has a contract @{term "x"} in @{term "X"}, then all of their contracts in @{const "XH_largest"} are at least as desirable as @{term "x"}, and so the @{const "stable_no_blocking"} hypothesis guarantees the hospitals choose @{term "x"} from @{const "XH_largest"}, and similarly the doctors @{term "x"} from @{const "XD_smallest"}. \ lemma XH_largestCdXXH_largest: assumes "x \ Ch h XH_largest" shows "x \ Cd (Xd x) (X \ Ch h XH_largest)" proof - from assms have "(y, x) \ Pd (Xd x)" if "Xd y = Xd x" and "y \ X" for y using that by (fastforce simp: XH_largest_def dX_def dest: Ch_range') with Ch_XH_largest_Field[OF assms] Pd_linear Pd_range show ?thesis using assms Ch_XH_largest_Field[OF assms] by (clarsimp simp: Cd_greatest greatest_def) (metis Ch_singular Pd_range' inj_onD subset_refl underS_incl_iff) qed lemma CH_XH_largest: shows "CH XH_largest = X" proof - have "Ch h XH_largest \ CD_on ds (X \ Ch h XH_largest)" for h using XH_largestCdXXH_largest Ch_XH_largest_Xd Ch_XH_largest_Field unfolding CD_on_def by blast from \stable_on ds X\ have "Ch h XH_largest = Ch h X" for h using \Ch h XH_largest \ CD_on ds (X \ Ch h XH_largest)\ X_subseteq_XH_largest by - (erule stable_on_blocking_onD[where h=h and X''="Ch h XH_largest"]; force intro!: consistencyD[OF Ch_consistency] dest: Ch_range') with stable_on_CH[OF \stable_on ds X\] show ?thesis unfolding CH_def by simp qed lemma Cd_XD_smallest: assumes "d \ ds" shows "Cd d (XD_smallest \ Field (Pd d)) = Cd d (X \ Field (Pd d))" proof(cases "X \ Field (Pd d) = {}") case True with Pd_range' Cd_range'[where X=X] stable_on_CD_on[OF \stable_on ds X\] mem_CD_on_Cd assms have "- XH_largest \ Field (Pd d) = {}" unfolding XH_largest_def dX_def by auto blast then show ?thesis unfolding XD_smallest_def by (simp add: Int_Un_distrib2) next case False with Pd_linear'[of d] \stable_on ds X\ stable_on_CD_on stable_on_allocation assms show ?thesis unfolding XD_smallest_def order_on_defs total_on_def by (auto 0 0 simp: Int_Un_distrib2 Cd_greatest greatest_def XH_largest_def dX_def) (metis (mono_tags, lifting) IntI Pd_range' UnCI inj_onD)+ qed lemma CD_on_XD_smallest: shows "CD_on ds XD_smallest = X" using stable_on_CD_on[OF \stable_on ds X\] unfolding CD_on_def2 by (simp add: Cd_XD_smallest) theorem stable_on_stable_pair_on: shows "stable_pair_on ds (XD_smallest, XH_largest)" proof(rule stable_pair_onI, simp_all only: prod.sel) from CH_XH_largest have "- RH XH_largest = - (XH_largest - X)" by blast also from X_XD_smallest_XH_largest have "\ = XD_smallest" unfolding XD_smallest_def by blast finally show "XD_smallest = -RH XH_largest" by blast next from CD_on_XD_smallest have "-RD_on ds XD_smallest = -(XD_smallest - X)" by simp also have "\ = XH_largest" unfolding XD_smallest_def using X_subseteq_XH_largest by blast finally show "XH_largest = -RD_on ds XD_smallest" by blast qed end end text\ Our ultimate statement of Theorem~1 of \cite{HatfieldMilgrom:2005} ala \citet{AygunSonmez:2012-WP2} goes as follows, bearing in mind that we are working in the @{const "ContractsWithIRC"} locale: \ theorem T1: shows "stable_on ds X \ (\XD_XH. stable_pair_on ds XD_XH \ X = match XD_XH)" using stable_pair_on_stable_on stable_on_stable_pair_on X_XD_smallest_XH_largest by fastforce end subsection\ Theorem~3: Algorithmics \label{sec:contracts-algorithmics} \ text (in Contracts) \ Having revived Theorem~1, we reformulate @{const "stable_pair"} as a monotone (aka @{emph \isotone\}) function and exploit the lattice structure of its fixed points, following \citet[{\S}II, Theorem~3]{HatfieldMilgrom:2005}. This underpins all of their results that we formulate here. \citet[\S2]{Fleiner:2002} provides an intuitive gloss of these definitions. \ context Contracts begin definition F1 :: "'x cfun" where "F1 X' = - RH X'" definition F2 :: "'d set \ 'x cfun" where "F2 ds X' = - RD_on ds X'" definition F :: "'d set \ 'x set \ 'x set dual \ 'x set \ 'x set dual" where "F ds = (\(XD, XH). (F1 (undual XH), dual (F2 ds (F1 (undual XH)))))" text\ We exploit Isabelle/HOL's ordering type classes (over the type constructors @{typ "'a set"} and @{typ "'a \ 'b"}) to define @{const "F"}. As @{const "F"} is @{const "antimono"} (where @{thm "antimono_def"} for a lattice order \\\) on its second argument \XH\, we adopt the dual lattice order using the type constructor @{typ "'a dual"}, where @{const "dual"} and @{const "undual"} mediate the isomorphism on values, to satisfy Isabelle/HOL's @{const "mono"} predicate. Note we work under the @{const "substitutes"} hypothesis here. Relating this function to @{const "stable_pair"} is syntactically awkward but straightforward: \ lemma fix_F_stable_pair_on: assumes "X = F ds X" shows "stable_pair_on ds (map_prod id undual X)" using %invisible assms by (cases X) (simp add: F_def F1_def F2_def stable_pair_on_def dual_eq_iff) lemma stable_pair_on_fix_F: assumes "stable_pair_on ds X" shows "map_prod id dual X = F ds (map_prod id dual X)" using %invisible assms unfolding F_def F1_def F2_def stable_pair_on_def split_def by (metis fst_map_prod id_apply prod.collapse snd_map_prod undual_dual) end text (in Contracts) \ The function @{const F} is monotonic under @{const substitutes}. \ locale ContractsWithSubstitutes = Contracts + assumes Ch_substitutes: "\h. substitutes (Ch h)" begin (*<*) lemma Rh_mono: shows "mono (Rh h)" using %invisible substitutes_on_Rf_mono_on[OF spec[OF Ch_substitutes]] mono_on_mono by (simp add: fun_eq_iff) blast lemmas Ch_iia = Rh_mono[unfolded Rf_mono_iia] lemmas Ch_Chernoff = Ch_iia[unfolded Chernoff_on_iia_on[symmetric]] lemmas Ch_subsitutes_idem = iia_f_idem[OF Ch_f_range Ch_iia, simplified] lemma RH_mono: shows "mono RH" unfolding %invisible CH_def by (rule monoI) (auto dest: monoD[OF Rh_mono]) lemmas CH_iia = RH_mono[unfolded Rf_mono_iia] lemmas CH_Chernoff = CH_iia[unfolded Chernoff_on_iia_on[symmetric]] lemmas CH_substitutes_idem = iia_f_idem[OF CH_f_range_on CH_iia, simplified] (*>*) text\\ lemma F1_antimono: shows "antimono F1" by %invisible (rule antimonoI) (auto simp: F1_def dest: Diff_mono[OF _ monoD[OF RH_mono]]) lemma F2_antimono: shows "antimono (F2 ds)" by %invisible (rule antimonoI) (auto simp: F2_def dest: Diff_mono[OF _ monoD[OF RD_on_mono]]) lemma F_mono: shows "mono (F ds)" unfolding %invisible F_def using antimonoD[OF F1_antimono] antimonoD[OF F2_antimono] by (auto intro: monoI simp: less_eq_dual_def) text\ We define the extremal fixed points using Isabelle/HOL's least and greatest fixed point operators: \ definition gfp_F :: "'b set \ 'a set \ 'a set" where "gfp_F ds = map_prod id undual (gfp (F ds))" definition lfp_F :: "'b set \ 'a set \ 'a set" where "lfp_F ds = map_prod id undual (lfp (F ds))" lemmas gfp_F_stable_pair_on = fix_F_stable_pair_on[OF gfp_unfold[OF F_mono], folded gfp_F_def] lemmas lfp_F_stable_pair_on = fix_F_stable_pair_on[OF lfp_unfold[OF F_mono], folded lfp_F_def] text\ These last two lemmas show that the least and greatest fixed points do satisfy @{const "stable_pair"}. Using standard fixed-point properties, we can establish: \ lemma F2_o_F1_mono: shows "mono (F2 ds \ F1)" by %invisible (metis F2_antimono F1_antimono antimono_def comp_apply monoI) lemmas F2_F1_mono = F2_o_F1_mono[unfolded o_def] lemma gfp_F_lfp_F: shows "gfp_F ds = (F1 (lfp (F2 ds \ F1)), lfp (F2 ds \ F1))" proof %invisible - let ?F' = "dual \ F2 ds \ F1 \ undual" have "gfp (F ds) = (F1 (undual (gfp ?F')), gfp ?F')" by (subst gfp_prod[OF F_mono]) (simp add: F_def o_def gfp_const) also have "gfp ?F' = dual (lfp (F2 ds \ F1))" by (simp add: lfp_dual_gfp[OF F2_o_F1_mono, simplified o_assoc]) finally show ?thesis unfolding gfp_F_def by simp qed end text\ We need hospital CFs to satisfy both @{const substitutes} and @{const irc} to relate these fixed points to stable matches. \ locale ContractsWithSubstitutesAndIRC = ContractsWithSubstitutes + ContractsWithIRC begin lemmas gfp_F_stable_on = stable_pair_on_stable_on[OF gfp_F_stable_pair_on] lemmas lfp_F_stable_on = stable_pair_on_stable_on[OF lfp_F_stable_pair_on] end text\ \label{sec:contracts-codegen-gfp_F} We demonstrate the effectiveness of our definitions by executing an example due to \citet[p920]{HatfieldMilgrom:2005} using Isabelle/HOL's code generator \citep{Haftmann-Nipkow:2010:code}. Note that, while adequate for this toy instance, the representations used here are hopelessly n{\"a}ive: sets are represented by lists and operations typically traverse the entire contract space. It is feasible, with more effort, to derive efficient algorithms; see, for instance, \citet{Bijlsma:1991,Bulwahn-et-al:2008:imp_HOL}. \ context ContractsWithSubstitutes begin lemma gfp_F_code[code]: shows "gfp_F ds = map_prod id undual (while (\A. F ds A \ A) (F ds) top)" using %invisible gfp_F_def gfp_while_lattice[OF F_mono] by simp lemma lfp_F_code[code]: shows "lfp_F ds = map_prod id undual (while (\A. F ds A \ A) (F ds) bot)" using %invisible lfp_F_def lfp_while_lattice[OF F_mono] by simp end text\ There are two hospitals and two doctors. \ datatype H2 = H1 | H2 text\ The contract space is simply the Cartesian product @{typ "D2 \ H2"}. \ type_synonym X_D2_H2 = "D2 \ H2" text\ Doctor @{const "D1"} prefers \H1 \ H2\, doctor @{const "D2"} the same \H1 \ H2\ (but over different contracts). \ primrec P_D2_H2_d :: "D2 \ X_D2_H2 rel" where "P_D2_H2_d D1 = linord_of_list [(D1, H1), (D1, H2)]" | "P_D2_H2_d D2 = linord_of_list [(D2, H1), (D2, H2)]" text\ Hospital @{const "H1"} prefers \{D1} \ {D2} \ \\, and hospital @{const "H2"} \{D1, D2} \ {D1} \ {D2} \ \\. We interpret these constraints as follows: \ definition P_D2_H2_H1 :: "X_D2_H2 cfun" where "P_D2_H2_H1 A = (if (D1, H1) \ A then {(D1, H1)} else if (D2, H1) \ A then {(D2, H1)} else {})" definition P_D2_H2_H2 :: "X_D2_H2 cfun" where "P_D2_H2_H2 A = (if {(D1, H2), (D2, H2)} \ A then {(D1, H2), (D2, H2)} else if (D1, H2) \ A then {(D1, H2)} else if (D2, H2) \ A then {(D2, H2)} else {})" primrec P_D2_H2_h :: "H2 \ X_D2_H2 cfun" where "P_D2_H2_h H1 = P_D2_H2_H1" | "P_D2_H2_h H2 = P_D2_H2_H2" (*<*) lemma H2_UNIV: shows "UNIV = set [H1, H2]" using H2.exhaust by auto instantiation H2 :: enum begin definition "enum_class.enum = [H1, H2]" definition "enum_class.enum_all P = (P H1 \ P H2)" definition "enum_class.enum_ex P = (P H1 \ P H2)" instance by standard (simp_all add: enum_H2_def enum_all_H2_def enum_ex_H2_def H2_UNIV) end lemma H2_ALL [simp]: shows "(\h. P h) = (\h\{H1, H2}. P h)" using H2_UNIV by auto lemma H2_UNION: shows "(\h. P h) = (\h\{H1, H2}. P h)" using H2_UNIV by auto lemma P_D2_H2_d_linear: shows "Linear_order (P_D2_H2_d d)" by (cases d) (simp_all add: linord_of_list_Linear_order) lemma P_D2_H2_d_range: shows "Field (P_D2_H2_d d) \ {x. fst x = d}" by (cases d) simp_all lemma P_D2_H2_h_substitutes: shows "substitutes (P_D2_H2_h h)" by %invisible (cases h) (auto intro!: substitutes_onI simp: P_D2_H2_H1_def P_D2_H2_H2_def split: if_splits) (*>*) text\ Isabelle's code generator requires us to hoist the relevant definitions from the locale to the top-level (see the \verb!codegen! documentation, \S7.3). \ global_interpretation P920_example: ContractsWithSubstitutes fst snd P_D2_H2_d P_D2_H2_h defines P920_example_gfp_F = P920_example.gfp_F and P920_example_lfp_F = P920_example.lfp_F and P920_example_F = P920_example.F and P920_example_F1 = P920_example.F1 and P920_example_F2 = P920_example.F2 and P920_example_maxR = P920_example.maxR and P920_example_MaxR_f = P920_example.MaxR_f and P920_example_Cd = P920_example.Cd and P920_example_CD_on = P920_example.CD_on and P920_example_CH = P920_example.CH using %invisible P_D2_H2_d_linear P_D2_H2_h_substitutes by %invisible unfold_locales (simp_all, simp_all add: D2_ALL P_D2_H2_H1_def P_D2_H2_H2_def) (*<*) (* Codegen hackery: avoid the CoSet constructor as some operations do not handle it. *) declare UNIV_coset[code del] declare UNIV_enum[code] declare compl_set[code del] compl_coset[code del] declare Compl_eq_Diff_UNIV[code] (* code_thms P920_example_gfp_F export_code P920_example_gfp_F in SML module_name F file "F.sml" value "P920_example_gfp_F UNIV" *) lemma P920_example_gfp_F_value: shows "P920_example_gfp_F UNIV = ({(D1, H1), (D1, H2), (D2, H2)}, {(D1, H1), (D2, H1), (D2, H2)})" by eval lemma P920_example_gfp_F_match_value: shows "P920_example.match (P920_example_gfp_F UNIV) = {(D1, H1), (D2, H2)}" unfolding P920_example_gfp_F_value by simp lemma P920_example_lfp_F_value: shows "P920_example_lfp_F UNIV = ({(D1, H1), (D1, H2), (D2, H2)}, {(D1, H1), (D2, H1), (D2, H2)})" by eval (*>*) text\ We can now evaluate the @{const "gfp"} of @{const "P920_example.F"} (i.e., \F\ specialized to the above constants) using Isabelle's \verb!value! antiquotation or \verb!eval! method. This yields the \(XD, XH)\ pair: \begin{center} @{thm (rhs) "P920_example_gfp_F_value"} \end{center} The stable match is therefore @{thm (rhs) "P920_example_gfp_F_match_value"}. The @{const "lfp"} of @{const "P920_example.F"} is identical to the @{const "gfp"}: \begin{center} @{thm (rhs) "P920_example_lfp_F_value"} \end{center} This implies that there is only one stable match in this scenario. \ subsection\ Theorem~4: Optimality \label{sec:contracts-optimality} \ text (in ContractsWithSubstitutes) \ \citet[Theorem~4]{HatfieldMilgrom:2005} assert that the greatest fixed point @{const "gfp_F"} of @{const "F"} yields the stable match most preferred by the doctors in the following sense: \ context Contracts begin definition doctor_optimal_match :: "'d set \ 'x set \ bool" where "doctor_optimal_match ds Y \ (stable_on ds Y \ (\X. \x\X. stable_on ds X \ (\y \ Y. (x, y) \ Pd (Xd x))))" (*<*) lemmas doctor_optimal_matchI = iffD2[OF doctor_optimal_match_def, unfolded conj_imp_eq_imp_imp, rule_format] lemmas doctor_optimal_match_stable_on = iffD1[OF doctor_optimal_match_def, THEN conjunct1] lemmas doctor_optimal_match_optimal = iffD1[OF doctor_optimal_match_def, THEN conjunct2, rule_format] lemma doctor_optimal_match_unique: assumes "doctor_optimal_match ds X" assumes "doctor_optimal_match ds Y" shows "X = Y" proof(rule iffD2[OF set_eq_iff, rule_format]) fix x from Pd_linear'[where d="Xd x"] Pd_Xd[where d="Xd x"] stable_on_allocation[OF doctor_optimal_match_stable_on[OF assms(1)]] stable_on_allocation[OF doctor_optimal_match_stable_on[OF assms(2)]] assms show "x \ X \ x \ Y" unfolding doctor_optimal_match_def order_on_defs by - (rule iffI; metis antisymD inj_on_eq_iff) qed (*>*) end text (in ContractsWithSubstitutes) \ In a similar sense, @{const "lfp_F"} is the doctor-pessimal match. We state a basic doctor-optimality result in terms of @{const "stable_pair"} in the @{const "ContractsWithSubstitutes"} locale for generality; we can establish @{const "doctor_optimal_match"} only under additional constraints on hospital choice functions (see \S\ref{sec:contracts-irc}). \ context ContractsWithSubstitutes begin context fixes XD_XH :: "'a set \ 'a set" fixes ds :: "'b set" assumes "stable_pair_on ds XD_XH" begin lemma gfp_F_upperbound: shows "(fst XD_XH, dual (snd XD_XH)) \ gfp (F ds)" proof %invisible - have "(fst XD_XH, dual (snd XD_XH)) = F ds (fst XD_XH, dual (snd XD_XH))" using stable_pair_on_fix_F[OF \stable_pair_on ds XD_XH\] by (metis id_apply map_prod_simp prod.collapse) then show ?thesis by (fastforce intro: gfp_upperbound) qed lemma XD_XH_gfp_F: shows "fst XD_XH \ fst (gfp_F ds)" and "snd (gfp_F ds) \ snd XD_XH" using %invisible gfp_F_upperbound unfolding gfp_F_def by (simp_all add: less_eq_dual_def less_eq_prod_def) lemma lfp_F_upperbound: shows "lfp (F ds) \ (fst XD_XH, dual (snd XD_XH))" proof %invisible - have "(fst XD_XH, dual (snd XD_XH)) = F ds (fst XD_XH, dual (snd XD_XH))" using stable_pair_on_fix_F[OF \stable_pair_on ds XD_XH\] by (metis id_apply map_prod_simp prod.collapse) then show ?thesis by (fastforce intro: lfp_lowerbound) qed lemma XD_XH_lfp_F: shows "fst (lfp_F ds) \ fst XD_XH" and "snd XD_XH \ snd (lfp_F ds)" using %invisible lfp_F_upperbound unfolding lfp_F_def by (simp_all add: less_eq_dual_def less_eq_prod_def) text\ We appeal to the doctors' linear preferences to show the optimality (pessimality) of @{const "gfp_F"} (@{const "lfp_F"}) for doctors. \ theorem gfp_f_doctor_optimal: assumes "x \ match XD_XH" shows "\y \ match (gfp_F ds). (x, y) \ Pd (Xd x)" using %invisible assms gfp_F_stable_pair_on[where ds=ds] \stable_pair_on ds XD_XH\ stable_pair_on_CD_on stable_pair_on_Xd Cd_Xd mem_CD_on_Cd XD_XH_gfp_F(1) Cd_mono[where d="Xd x" and x=x and X="fst XD_XH" and Y="fst (gfp_F ds)"] by (metis sup.absorb_iff2) theorem lfp_f_doctor_pessimal: assumes "x \ match (lfp_F ds)" shows "\y \ match XD_XH. (x, y) \ Pd (Xd x)" using %invisible assms lfp_F_stable_pair_on[where ds=ds] \stable_pair_on ds XD_XH\ stable_pair_on_CD_on stable_pair_on_Xd Cd_Xd mem_CD_on_Cd XD_XH_lfp_F(1) Cd_mono[where d="Xd x" and x=x and X="fst (lfp_F ds)" and Y="fst XD_XH"] by (metis sup.absorb_iff2) end end theorem (in ContractsWithSubstitutesAndIRC) gfp_F_doctor_optimal_match: shows "doctor_optimal_match ds (match (gfp_F ds))" by %invisible (rule doctor_optimal_matchI[OF gfp_F_stable_on]) (auto simp: T1 elim: gfp_f_doctor_optimal) text (in ContractsWithSubstitutesAndIRC) \ Conversely @{const "lfp_F"} is most preferred by the hospitals in a revealed-preference sense, and @{const "gfp_F"} least preferred. These results depend on @{thm [source] Ch_domain} and hence the @{const "irc"} hypothesis on hospital choice functions. \ context ContractsWithSubstitutesAndIRC begin theorem lfp_f_hospital_optimal: assumes "stable_pair_on ds XD_XH" assumes "x \ Ch h (match (lfp_F ds))" shows "x \ Ch h (match (lfp_F ds) \ match XD_XH)" proof %invisible - from \stable_pair_on ds XD_XH\ have "match (lfp_F ds) \ match XD_XH \ snd (lfp_F ds)" by (simp add: XD_XH_lfp_F(2) le_infI2) with \x \ Ch h (match (lfp_F ds))\ lfp_F_stable_pair_on stable_pair_on_Ch Ch_range show ?thesis by - (rule iia_onD[OF Ch_iia[where h=h], where B="snd (lfp_F ds)", simplified]; blast) qed theorem gfp_f_hospital_pessimal: assumes "stable_pair_on ds XD_XH" assumes "x \ Ch h (match XD_XH)" shows "x \ Ch h (match (gfp_F ds) \ match XD_XH)" proof %invisible - from \stable_pair_on ds XD_XH\ have "match (gfp_F ds) \ match XD_XH \ snd XD_XH" by (simp add: XD_XH_gfp_F(2) le_infI2) with assms lfp_F_stable_pair_on stable_pair_on_Ch Ch_range show ?thesis by - (rule iia_onD[OF Ch_iia[where h=h], where B="snd XD_XH", simplified]; blast+) qed end text\ The general lattice-theoretic results of e.g. \citet{Fleiner:2002} depend on the full Tarski-Knaster fixed point theorem, which is difficult to state in the present type class-based setting. (The theorem itself is available in the Isabelle/HOL distribution but requires working with less convenient machinery.) \ subsection\ Theorem~5 does not hold \citep{HatfieldKojima:2008} \ text (in Contracts) \ \citet[Theorem~5]{HatfieldMilgrom:2005} claim that: \begin{quote} Suppose \H\ contains at least two hospitals, which we denote by \h\ and \h'\. Further suppose that @{term "Rh h"} is not isotone, that is, contracts are not @{const "substitutes"} for \h\. Then there exist preference orderings for the doctors in set \D\, a preference ordering for a hospital \h'\ with a single job opening such that, regardless of the preferences of the other hospitals, no stable set of contracts exists. \end{quote} \citet[Observation~1]{HatfieldKojima:2008} show this is not true: there can be stable matches even if hospital choice functions violate @{const "substitutes"}. This motivates looking for conditions weaker than @{const "substitutes"} that still guarantee stable matches, a project taken up by \citet{HatfieldKojima:2010}; see \S\ref{sec:cop}. We omit their counterexample to this incorrect claim. \ subsection\ Theorem~6: ``Vacancy chain'' dynamics \label{sec:contracts-vacancy-chain} \ text (in ContractsWithSubstitutesAndIRC) \ \citet[II(C), p923]{HatfieldMilgrom:2005} propose a model for updating a stable match @{term "X"} when a doctor @{term "d'"} retires. Intuitively the contracts mentioning @{term "d'"} are discarded and a modified algorithm run from the @{const "XH_largest"} and @{const "XD_smallest"} sets determined from @{term "X"}. The result is another stable match where the remaining doctors @{term "ds - {d'}"} are (weakly) better off and the hospitals (weakly) worse off than they were in the initial state. The proofs are essentially the same as for optimality (\S\ref{sec:contracts-optimality}). \ context ContractsWithSubstitutesAndIRC begin context fixes X :: "'a set" fixes d' :: "'b" fixes ds :: "'b set" begin text\ \citeauthor{HatfieldMilgrom:2005} do not motivate why the process uses this functional and not @{const "F"}. \ definition F' :: "'a set \ 'a set dual \ 'a set \ 'a set dual" where "F' = (\(XD, XH). (- RH (undual XH), dual (- RD_on (ds-{d'}) XD)))" lemma F'_apply: "F' (XD, XH) = (- RH (undual XH), dual (- RD_on (ds - {d'}) XD))" by (simp add: F'_def) lemma %invisible F1'_antimono: shows "antimono (\XH. - RH XH)" by %invisible (rule antimonoI) (auto simp: F1_def dest: Diff_mono[OF _ monoD[OF RH_mono]]) lemma %invisible F2'_antimono: shows "antimono (\XD. - RD_on (ds-{d'}) XD)" by %invisible (rule antimonoI) (auto simp: F2_def dest: Diff_mono[OF _ monoD[OF RD_on_mono]]) lemma F'_mono: shows "mono F'" unfolding %invisible F'_def using antimonoD[OF F1'_antimono] antimonoD[OF F2'_antimono] by (auto intro: monoI simp: less_eq_dual_def) lemma fix_F'_stable_pair_on: "stable_pair_on (ds - {d'}) (map_prod id undual A)" if "A = F' A" proof %invisible - obtain x y where "A = (x, y)" by (cases A) with that have "F' (x, y) = (x, y)" by simp then have "- Rf CH (undual y) = x" and "dual (- Rf (CD_on (ds - {d'})) x) = y" by (simp_all only: F'_apply prod_eq_iff fst_conv snd_conv) with \A = (x, y)\ show ?thesis by (simp add: stable_pair_on_def dual_eq_iff) qed text\ We model their update process using the @{const "while"} combinator, as we cannot connect it to the extremal fixed points as we did in \S\ref{sec:contracts-algorithmics} because we begin computing from the stable match @{term "X"}. \ definition F'_iter :: "'a set \ 'a set dual" where "F'_iter = (while (\A. F' A \ A) F' (XD_smallest ds X, dual (XH_largest ds X)))" abbreviation F'_iter_match :: "'a set" where "F'_iter_match \ match (map_prod id undual F'_iter)" context assumes "stable_on ds X" begin lemma F_start: shows "F ds (XD_smallest ds X, dual (XH_largest ds X)) = (XD_smallest ds X, dual (XH_largest ds X))" using %invisible CH_XH_largest[OF \stable_on ds X\] CD_on_XD_smallest[OF \stable_on ds X\] X_subseteq_XH_largest[OF \stable_on ds X\] unfolding F_def F1_def F2_def XD_smallest_def by (auto simp add: dual_eq_iff) lemma F'_start: shows "(XD_smallest ds X, dual (XH_largest ds X)) \ F' (XD_smallest ds X, dual (XH_largest ds X))" using %invisible F_start unfolding F_def F1_def F2_def F'_def unfolding CD_on_def XD_smallest_def by (auto simp add: dual_eq_iff dual_less_eq_iff) lemma shows F'_iter_stable_pair_on: "stable_pair_on (ds-{d'}) (map_prod id undual F'_iter)" (is "?thesis1") and F'_start_le_F'_iter: "(XD_smallest ds X, dual (XH_largest ds X)) \ F'_iter" (is "?thesis2") proof %invisible - obtain P where XXX: "while_option (\A. F' A \ A) F' ((XD_smallest ds X), dual (XH_largest ds X)) = Some P" using while_option_finite_increasing_Some[OF F'_mono _ F'_start, simplified] by blast with while_option_stop2[OF XXX] fix_F'_stable_pair_on[where A=P] show ?thesis1 and ?thesis2 using funpow_mono2[OF F'_mono _ order.refl F'_start, where i=0] unfolding F'_iter_def while_def by auto qed lemma F'_iter_match_stable_on: shows "stable_on (ds-{d'}) F'_iter_match" by %invisible (rule stable_pair_on_stable_on) (metis F'_iter_stable_pair_on) theorem F'_iter_match_doctors_weakly_better_off: assumes "x \ Cd d X" assumes "d \ d'" shows "\y \ Cd d F'_iter_match. (x, y) \ Pd d" proof %invisible - from \stable_on ds X\ assms have "d \ ds" by (blast dest: Cd_Xd Cd_range' stable_on_Xd) with assms \stable_on ds X\ stable_on_stable_pair_on[OF \stable_on ds X\] have "\y\Cd d (XD_smallest ds X \ fst F'_iter). (x, y) \ Pd d" by - (rule Cd_mono; fastforce dest: X_XD_smallest_XH_largest stable_pair_on_Cd_match) with F'_iter_stable_pair_on F'_start_le_F'_iter have "\y\Cd d (fst F'_iter). (x, y) \ Pd d" by (metis Pair_le Un_absorb1 prod.collapse) with \d \ d'\ \d \ ds\ show ?thesis using stable_pair_on_Cd[OF F'_iter_stable_pair_on, symmetric, of d] by (subst Cd_domain[symmetric]) (simp add: Cd_idem) qed theorem F'_iter_match_hospitals_weakly_worse_off: assumes "x \ Ch h X" shows "x \ Ch h (F'_iter_match \ X)" proof %invisible - from F'_iter_stable_pair_on F'_start_le_F'_iter stable_on_stable_pair_on[OF \stable_on ds X\] X_subseteq_XH_largest[OF \stable_on ds X\] have "F'_iter_match \ X \ XH_largest ds X" by (simp add: less_eq_prod_def le_infI2 less_eq_dual_def) with assms Ch_range \stable_on ds X\ show ?thesis by - (rule iia_onD[OF Ch_iia, where B="XH_largest ds X"], auto, metis Ch_CH_irc_idem CH_XH_largest) qed text\ \citeauthor{HatfieldMilgrom:2005} observe but do not prove that @{const "F'_iter_match"} is not necessarily doctor-optimal wrt the new set of doctors, even if @{term "X"} was. These results seem incomplete. One might expect that the process of reacting to a doctor's retirement would involve considering new entrants to the workforce and allowing the set of possible contracts to be refined. There are also the questions of hospitals opening and closing. \ end end end subsection\ Theorems~8~and~9: A ``rural hospitals'' theorem \label{sec:contracts-rh} \ text\ Given that some hospitals are less desirable than others, the question arises of whether there is a mechanism that can redistribute doctors to under-resourced hospitals while retaining the stability of the match. Roth's @{emph \rural hospitals theorem\} \citep[Theorem~5.12]{RothSotomayor:1990} resolves this in the negative by showing that each doctor and hospital signs the same number of contracts in every stable match. In the context of contracts the theorem relies on the further hypothesis that hospital choices satisfy the law of aggregate demand (\S\ref{sec:cf-lad}). \ locale ContractsWithLAD = Contracts + assumes Ch_lad: "\h. lad (Ch h)" locale ContractsWithSubstitutesAndLAD = ContractsWithSubstitutes + ContractsWithLAD text\ We can use results that hold under @{const "irc"} by discharging that hypothesis against @{const "lad"} using the @{thm [source] "lad_on_substitutes_on_irc_on"} lemma. This is the effect of the following \sublocale\ command: \ sublocale ContractsWithSubstitutesAndLAD < ContractsWithSubstitutesAndIRC using Ch_range Ch_substitutes Ch_lad by unfold_locales (blast intro: lad_on_substitutes_on_irc_on f_range_onI) context ContractsWithSubstitutesAndLAD begin text\ The following results lead to \citet[Theorem~8]{HatfieldMilgrom:2005}, and the proofs go as they say. Again we state these with respect to an arbitrary solution to @{const "stable_pair"}. \ context fixes XD_XH :: "'a set \ 'a set" fixes ds :: "'b set" assumes "stable_pair_on ds XD_XH" begin lemma Cd_XD_gfp_F_card: assumes "d \ ds" shows "card (Cd d (fst XD_XH)) \ card (Cd d (fst (gfp_F ds)))" using %invisible assms Cd_lad XD_XH_gfp_F(1)[OF \stable_pair_on ds XD_XH\] unfolding lad_on_def by blast lemma Ch_gfp_F_XH_card: shows "card (Ch h (snd (gfp_F ds))) \ card (Ch h (snd XD_XH))" using %invisible Ch_lad XD_XH_gfp_F(2)[OF \stable_pair_on ds XD_XH\] unfolding lad_on_def by blast theorem Theorem_8: shows "d \ ds \ card (Cd d (fst XD_XH)) = card (Cd d (fst (gfp_F ds)))" and "card (Ch h (snd XD_XH)) = card (Ch h (snd (gfp_F ds)))" proof %invisible - let ?Sum_Cd_gfp = "\d\ds. card (Cd d (fst (gfp_F ds)))" let ?Sum_Ch_gfp = "\h\UNIV. card (Ch h (snd (gfp_F ds)))" let ?Sum_Cd_XD = "\d\ds. card (Cd d (fst XD_XH))" let ?Sum_Ch_XH = "\h\UNIV. card (Ch h (snd XD_XH))" have "?Sum_Cd_gfp = ?Sum_Ch_gfp" using stable_pair_on_CD_on_CH[OF gfp_F_stable_pair_on] CD_on_card[symmetric] CH_card[symmetric] by simp also have "\ \ ?Sum_Ch_XH" using Ch_gfp_F_XH_card by (simp add: sum_mono) also have "\ = ?Sum_Cd_XD" using stable_pair_on_CD_on_CH[OF \stable_pair_on ds XD_XH\] CD_on_card[symmetric] CH_card[symmetric] by simp finally have "?Sum_Cd_XD = ?Sum_Cd_gfp" using Cd_XD_gfp_F_card by (simp add: eq_iff sum_mono) with Cd_XD_gfp_F_card show "d \ ds \ card (Cd d (fst XD_XH)) = card (Cd d (fst (gfp_F ds)))" by (fastforce elim: sum_mono_inv) have "?Sum_Ch_XH = ?Sum_Cd_XD" using stable_pair_on_CD_on_CH[OF \stable_pair_on ds XD_XH\] CD_on_card[symmetric] CH_card[symmetric] by simp also have "\ \ ?Sum_Cd_gfp" using Cd_XD_gfp_F_card by (simp add: sum_mono) also have "\ = ?Sum_Ch_gfp" using stable_pair_on_CD_on_CH[OF gfp_F_stable_pair_on] CD_on_card[symmetric] CH_card[symmetric] by simp finally have "?Sum_Ch_gfp = ?Sum_Ch_XH" using Ch_gfp_F_XH_card by (simp add: eq_iff sum_mono) with Ch_gfp_F_XH_card show "card (Ch h (snd XD_XH)) = card (Ch h (snd (gfp_F ds)))" by (fastforce elim: sym[OF sum_mono_inv]) qed end text\ Their result may be more easily understood when phrased in terms of arbitrary stable matches: \ corollary rural_hospitals_theorem: assumes "stable_on ds X" assumes "stable_on ds Y" shows "d \ ds \ card (Cd d X) = card (Cd d Y)" and "card (Ch h X) = card (Ch h Y)" using %invisible assms T1[of ds X] T1[of ds Y] Theorem_8 stable_pair_on_Cd_match Ch_CH_irc_idem stable_pair_on_CH by force+ end text\ \citet[Theorem~9]{HatfieldMilgrom:2005} show that without @{const "lad"}, the rural hospitals theorem does not hold. Their proof does not seem to justify the theorem as stated (for instance, the contracts \x'\, \y'\ and \z'\ need not exist), and so we instead simply provide a counterexample (discovered by \verb!nitpick!) to the same effect. \ lemma (in ContractsWithSubstitutesAndIRC) Theorem_9_counterexample: assumes "stable_on ds Y" assumes "stable_on ds Z" shows "card (Ch h Y) = card (Ch h Z)" oops datatype X3 = Xd1 | Xd1' | Xd2 (*<*) lemma X3_UNIV: shows "UNIV = set [Xd1, Xd1', Xd2]" using X3.exhaust by auto lemmas X3_pow = subset_subseqs[OF subset_trans[OF subset_UNIV Set.equalityD1[OF X3_UNIV]]] instance X3 :: finite by standard (simp add: X3_UNIV) lemma X3_all_pow: shows "(\X''. P X'') \ (\X''\set ` set (subseqs [Xd1, Xd1', Xd2]). P X'')" using X3_pow by blast (*>*) primrec X3d :: "X3 \ D2" where "X3d Xd1 = D1" | "X3d Xd1' = D1" | "X3d Xd2 = D2" abbreviation X3h :: "X3 \ H1" where "X3h _ \ H" primrec PX3d :: "D2 \ X3 rel" where "PX3d D1 = linord_of_list [Xd1, Xd1']" | "PX3d D2 = linord_of_list [Xd2]" function CX3h :: "H1 \ X3 set \ X3 set" where "CX3h _ {Xd1} = {Xd1}" | "CX3h _ {Xd1'} = {Xd1'}" | "CX3h _ {Xd2} = {Xd2}" | "CX3h _ {Xd1, Xd1'} = {Xd1'}" | "CX3h _ {Xd1, Xd2} = {Xd1, Xd2}" | "CX3h _ {Xd1', Xd2} = {Xd1'}" | "CX3h _ {Xd1, Xd1', Xd2} = {Xd1'}" | "CX3h _ {} = {}" apply %invisible (case_tac x) apply (cut_tac X=b in X3_pow) apply auto done (*<*) termination by lexicographic_order lemma PX3d_linear: shows "Linear_order (PX3d d)" by (cases d) (simp_all add: linord_of_list_Linear_order) lemma PX3d_range: shows "Field (PX3d d) \ {x. X3d x = d}" by (cases d) simp_all lemma CX3h_range: shows "CX3h h X \ {x\X. X3h x = h}" by (cases "(h, X)" rule: CX3h.cases; simp; metis (mono_tags) H1.exhaust) lemma CX3h_singular: shows "inj_on X3d (CX3h h X)" by (cases "(h, X)" rule: CX3h.cases) auto lemma CX3h_substitutes: shows "substitutes (CX3h h)" apply (rule substitutes_onI) apply (cases h) apply (cut_tac X=B in X3_pow) apply (case_tac b) apply (case_tac [!] a) apply (auto simp: insert_commute) done lemma CX3h_irc: shows "irc (CX3h h)" apply (rule ircI) apply (cases h) apply (cut_tac X=B in X3_pow) apply (case_tac a) apply (auto simp: insert_commute) done (*>*) interpretation Theorem_9: ContractsWithSubstitutesAndIRC X3d X3h PX3d CX3h using %invisible PX3d_linear PX3d_range CX3h_range CX3h_singular CX3h_substitutes CX3h_irc by unfold_locales blast+ lemma Theorem_9_stable_Xd1': shows "Theorem_9.stable_on UNIV {Xd1'}" proof %invisible (rule Theorem_9.stable_onI) note image_cong_simp [cong del] note INF_cong_simp [cong] note SUP_cong_simp [cong] show "Theorem_9.individually_rational_on UNIV {Xd1'}" by (rule Theorem_9.individually_rational_onI) (simp_all add: D2_UNION Theorem_9.CD_on_def Theorem_9.CH_def) show "Theorem_9.stable_no_blocking_on UNIV {Xd1'}" apply (rule Theorem_9.stable_no_blocking_onI) apply (case_tac h) apply (cut_tac X=X'' in X3_pow) apply simp apply safe apply (simp_all add: insert_commute) done qed lemma Theorem_9_stable_Xd1_Xd2: shows "Theorem_9.stable_on UNIV {Xd1, Xd2}" proof %invisible (rule Theorem_9.stable_onI) note image_cong_simp [cong del] note INF_cong_simp [cong] note SUP_cong_simp [cong] show "Theorem_9.individually_rational_on UNIV {Xd1, Xd2}" by (rule Theorem_9.individually_rational_onI) (simp_all add: D2_UNION Theorem_9.CD_on_def Theorem_9.CH_def insert_commute) show "Theorem_9.stable_no_blocking_on UNIV {Xd1, Xd2}" apply (rule Theorem_9.stable_no_blocking_onI) apply (case_tac h) apply (cut_tac X=X'' in X3_pow) apply simp apply safe apply (simp_all add: D2_UNION Theorem_9.CD_on_def Theorem_9.maxR_def linord_of_list_linord_of_listP insert_commute) done qed text \ This violates the rural hospitals theorem: \ theorem shows "card (Theorem_9.CH {Xd1'}) \ card (Theorem_9.CH {Xd1, Xd2})" using %invisible Theorem_9_stable_Xd1' Theorem_9_stable_Xd1_Xd2 Theorem_9.stable_on_CH by simp text\ {\ldots}which is attributed to the failure of the hospitals' choice functions to satisfy @{const "lad"}: \ lemma CX3h_not_lad: shows "\lad (CX3h h)" unfolding %invisible lad_on_def apply (cases h) apply clarsimp apply (rule exI[where x="{Xd1, Xd1', Xd2}"]) apply (rule exI[where x="{Xd1, Xd2}"]) apply simp done text\ \citet{CiupanHatfieldKominers:2016} discuss an alternative approach to this result in a marriage market. \ subsection\ Theorems~15 and 16: Cumulative Offer Processes \label{sec:contracts-cop} \ text\ The goal of \citet[{\S}V]{HatfieldMilgrom:2005} is to connect this theory of contracts with matching to earlier work on auctions by the first of the authors, in particular by eliminating the @{const "substitutes"} hypothesis. They do so by defining a @{emph \cumulative offer process\} (COP): \ context Contracts begin definition cop_F_HM :: "'d set \ 'x set \ 'x set \ 'x set \ 'x set" where "cop_F_HM ds = (\(XD, XH). (- RH XH, XH \ CD_on ds (- RH XH)))" text\ Intuitively all of the doctors simultaneously offer their most preferred contracts that have yet to be rejected by the hospitals, and the hospitals choose amongst these and all that have been offered previously. Asking hospital choice functions to satisfy the @{const "substitutes"} condition effectively forces hospitals to consider only the contracts they have previously not rejected. This definition is neither monotonic nor increasing (i.e., it is not the case that @{term "\x. x \ cop_F_HM ds x"}). We rectify this by focusing on the second part of the definition. \ definition cop_F :: "'d set \ 'x set \ 'x set" where "cop_F ds XH = XH \ CD_on ds (- RH XH)" lemma cop_F_HM_cop_F: shows "cop_F_HM ds XD_XH = (- RH (snd XD_XH), cop_F ds (snd XD_XH))" unfolding cop_F_HM_def cop_F_def split_def by simp lemma cop_F_increasing: shows "x \ cop_F ds x" unfolding %invisible cop_F_def by simp text\ We have the following straightforward case distinction principles: \ lemma cop_F_cases: assumes "x \ cop_F ds fp" obtains (fp) "x \ fp" | (CD_on) "x \ CD_on ds (-RH fp) - fp" using assms unfolding cop_F_def by blast lemma CH_cop_F_cases: assumes "x \ CH (cop_F ds fp)" obtains (CH) "x \ CH fp" | (RH_fp) "x \ RH fp" | (CD_on) "x \ CD_on ds (-RH fp) - fp" using assms CH_range cop_F_def by auto text\ The existence of fixed points for our earlier definitions (\S\ref{sec:contracts-algorithmics}) was guaranteed by the Tarski-Knaster theorem, which relies on the monotonicity of the defining functional. As @{const "cop_F"} lacks this property, we appeal instead to the Bourbaki-Witt theorem for increasing functions. \ interpretation COP: bourbaki_witt_fixpoint Sup "{(x, y). x \ y}" "cop_F ds" for ds by %invisible (rule bourbaki_witt_fixpoint_complete_latticeI[OF cop_F_increasing]) definition fp_cop_F :: "'d set \ 'x set" where "fp_cop_F ds = COP.fixp_above ds {}" abbreviation "cop ds \ CH (fp_cop_F ds)" (*<*) lemmas fp_cop_F_unfold = COP.fixp_above_unfold[where a="{}", folded fp_cop_F_def, simplified Field_def, simplified] lemmas fp_cop_F_code = COP.fixp_above_conv_while[where a="{}", folded fp_cop_F_def, simplified Field_def, simplified] (*>*) text\ Given that the set of contracts is finite, we avoid continuity and admissibility issues; we have the following straightforward induction principle: \ lemma fp_cop_F_induct[case_names base step]: assumes "P {}" assumes "\fp. P fp \ P (cop_F ds fp)" shows "P (fp_cop_F ds)" using %invisible assms by (induct rule: COP.fixp_above_induct[where a="{}", folded fp_cop_F_def]) (fastforce intro: admissible_chfin)+ text\ An alternative is to use the @{const "while"} combinator, which is equivalent to the above by @{thm [source] COP.fixp_above_conv_while}. In any case, invariant reasoning is essential to verifying the properties of the COP, no matter how we phrase it. We develop a small program logic to ease the reuse of the invariants we prove. \ definition valid :: "'d set \ ('d set \ 'x set \ bool) \ ('d set \ 'x set \ bool) \ bool" where "valid ds P Q = (Q ds {} \ (\fp. P ds fp \ Q ds fp \ Q ds (cop_F ds fp)))" abbreviation invariant :: "'d set \ ('d set \ 'x set \ bool) \ bool" where "invariant ds P \ valid ds (\_ _. True) P" text\ Intuitively @{term "valid ds P Q"} asserts that the COP satisfies @{term "Q"} assuming that it satisfies @{term "P"}. This allows us to decompose our invariant proofs. By setting the precondition to @{term "True"}, @{term "invariant ds P"} captures the proof obligations of @{term "fp_cop_F_induct"} exactly. The following lemmas ease the syntactic manipulation of these facts. \ lemma validI[case_names base step]: assumes "Q ds {}" assumes "\fp. \P ds fp; Q ds fp\ \ Q ds (cop_F ds fp)" shows "valid ds P Q" using %invisible assms unfolding valid_def by blast lemma invariant_cop_FD: assumes "invariant ds P" assumes "P ds fp" shows "P ds (cop_F ds fp)" using %invisible assms unfolding valid_def by blast lemma invariantD: assumes "invariant ds P" shows "P ds (fp_cop_F ds)" using %invisible assms fp_cop_F_induct unfolding valid_def by blast lemma valid_pre: assumes "valid ds P' Q" assumes "\fp. P ds fp \ P' ds fp" shows "valid ds P Q" using %invisible assms unfolding valid_def by blast lemma valid_invariant: assumes "valid ds P Q" assumes "invariant ds P" shows "invariant ds (\ ds fp. P ds fp \ Q ds fp)" using %invisible assms unfolding valid_def by blast lemma valid_conj: assumes "valid ds (\ds fp. R ds fp \ P ds fp \ Q ds fp) P" assumes "valid ds (\ds fp. R ds fp \ P ds fp \ Q ds fp) Q" shows "valid ds R (\ ds fp. P ds fp \ Q ds fp)" using %invisible assms unfolding valid_def by blast end text (in ContractsWithSubstitutes) \ \citet[Theorem~15]{HatfieldMilgrom:2005} assert that @{const "fp_cop_F"} is equivalent to the doctor-offering algorithm @{const "gfp_F"}, assuming @{const "substitutes"}. (Note that the fixed points generated by increasing functions do not necessarily form a lattice, so there is not necessarily a hospital-optimal match, and indeed in general these do not exist.) Our proof is eased by the decomposition lemma @{thm [source] gfp_F_lfp_F} and the standard properties of fixed points in a lattice. \ context ContractsWithSubstitutes begin lemma lfp_F2_o_F1_fp_cop_F: shows "lfp (F2 ds \ F1) = fp_cop_F ds" proof(rule antisym) have "(F2 ds \ F1) (fp_cop_F ds) \ cop_F ds (fp_cop_F ds)" by (clarsimp simp: F2_def F1_def cop_F_def) then show "lfp (F2 ds \ F1) \ fp_cop_F ds" by (simp add: lfp_lowerbound fp_cop_F_unfold[symmetric]) next show "fp_cop_F ds \ lfp (F2 ds \ F1)" proof(induct rule: fp_cop_F_induct) case base then show ?case by simp next case (step fp) note IH = \fp \ lfp (F2 ds \ F1)\ then have "CD_on ds (- RH fp) \ lfp (F2 ds \ F1)" by (subst lfp_unfold[OF F2_o_F1_mono]) (metis (no_types, lifting) Compl_Diff_eq F1_antimono F2_antimono F1_def F2_def Un_subset_iff antimonoD comp_apply) with IH show ?case unfolding cop_F_def by blast qed qed theorem Theorem_15: shows "gfp_F ds = (- RH (fp_cop_F ds), fp_cop_F ds)" using lfp_F2_o_F1_fp_cop_F unfolding gfp_F_lfp_F F1_def by simp theorem Theorem_15_match: shows "match (gfp_F ds) = CH (fp_cop_F ds)" using Theorem_15 by (fastforce dest: subsetD[OF CH_range]) end text\ \label{sec:contracts-codegen-fp_cop_F} With some auxiliary definitions, we can evaluate the COP on the example from \S\ref{sec:contracts-codegen-gfp_F}. \ (*<*) definition "P920_example_cop_F = P920_example.cop_F" definition "P920_example_fp_cop_F = P920_example.fp_cop_F" lemmas P920_example_cop_F_code[code] = P920_example.cop_F_def[folded P920_example_cop_F_def] lemmas P920_example_fp_cop_F_code[code] = P920_example.fp_cop_F_code[folded P920_example_fp_cop_F_def P920_example_cop_F_def] (*>*) lemma P920_example_fp_cop_F_value: shows "P920_example_CH (P920_example_fp_cop_F UNIV) = {(D1, H1), (D2, H2)}" by eval text\ \citet[Theorem~16]{HatfieldMilgrom:2005} assert that this process yields a stable match when we have a single hospital (now called an auctioneer) with unrestricted preferences. As before, this holds provided the auctioneer's preferences satisfy @{const "irc"}. We begin by establishing two obvious invariants of the COP that hold in general. \ context Contracts begin lemma %invisible CH_Ch_singular: assumes "(UNIV::'h set) = {h}" shows "CH A = Ch h A" unfolding CH_def using assms by auto definition cop_F_range_inv :: "'d set \ 'x set \ bool" where "cop_F_range_inv ds fp \ (\x\fp. x \ Field (Pd (Xd x)) \ Xd x \ ds)" definition cop_F_closed_inv :: "'d set \ 'x set \ bool" where "cop_F_closed_inv ds fp \ (\x\fp. above (Pd (Xd x)) x \ fp)" text\ The first, @{const "cop_F_range_inv"}, simply states that the result of the COP respects the structural conditions for doctors. The second @{const "cop_F_closed_inv"} states that the COP is upwards-closed with respect to the doctors' preferences. \ lemma cop_F_range_inv: shows "invariant ds cop_F_range_inv" unfolding valid_def cop_F_range_inv_def cop_F_def by (fastforce simp: mem_CD_on_Cd dest: Cd_range') lemma cop_F_closed_inv: shows "invariant ds cop_F_closed_inv" unfolding valid_def cop_F_closed_inv_def cop_F_def above_def by (clarsimp simp: subset_iff) (metis Cd_preferred ComplI Un_upper1 mem_CD_on_Cd subsetCE) lemmas fp_cop_F_range_inv = invariantD[OF cop_F_range_inv] lemmas fp_cop_F_range_inv' = fp_cop_F_range_inv[unfolded cop_F_range_inv_def, rule_format] lemmas fp_cop_F_closed_inv = invariantD[OF cop_F_closed_inv] lemmas fp_cop_F_closed_inv' = subsetD[OF bspec[OF invariantD[OF cop_F_closed_inv, unfolded cop_F_closed_inv_def, simplified]]] text\ The only challenge in showing that the COP yields a stable match is in establishing @{const "stable_no_blocking_on"}. Our key lemma states that that if @{const "CH"} rejects all contracts for doctor \d\ in @{const "fp_cop_F"}, then all contracts for \d\ are in @{const "fp_cop_F"}. \ lemma cop_F_RH: assumes "d \ ds" assumes "x \ Field (Pd d)" assumes "aboveS (Pd d) x \ RH fp" shows "x \ cop_F ds fp" using %invisible assms Pd_linear unfolding cop_F_def by (clarsimp simp: mem_CD_on_Cd Cd_greatest greatest_def aboveS_def order_on_defs total_on_def subset_iff) (metis Compl_Diff_eq Compl_iff Diff_iff IntE Pd_Xd refl_onD) lemma fp_cop_F_all: assumes "d \ ds" assumes "d \ Xd ` CH (fp_cop_F ds)" shows "Field (Pd d) \ fp_cop_F ds" proof %invisible (rule subsetI) fix x assume "x \ Field (Pd d)" from spec[OF Pd_linear] this finite[of "Pd d"] show "x \ fp_cop_F ds" proof(induct rule: finite_Linear_order_induct) case (step x) with assms Pd_range Pd_Xd cop_F_RH[of d ds _ "fp_cop_F ds", unfolded fp_cop_F_unfold[symmetric]] show ?case unfolding aboveS_def by (fastforce iff: image_iff) qed qed text\ \citet{AygunSonmez:2012-WP2} observe that any blocking contract must be weakly preferred by its doctor to anything in the outcome of the @{const "fp_cop_F"}: \ lemma fp_cop_F_preferred: assumes "y \ CD_on ds (CH (fp_cop_F ds) \ X'')" assumes "x \ CH (fp_cop_F ds)" assumes "Xd x = Xd y" shows "(x, y) \ Pd (Xd x)" using %invisible assms fp_cop_F_range_inv'[OF CH_range'[OF assms(2)]] Pd_Xd Pd_linear by (clarsimp simp: CD_on_def Cd_greatest greatest_def) (metis Int_iff Un_iff subset_refl underS_incl_iff) text\ The headline lemma cobbles these results together. \ lemma X''_closed: assumes "X'' \ CD_on ds (CH (fp_cop_F ds) \ X'')" shows "X'' \ fp_cop_F ds" proof(rule subsetI) fix x assume "x \ X''" show "x \ fp_cop_F ds" proof(cases "Xd x \ Xd ` CH (fp_cop_F ds)") case True then obtain y where "Xd y = Xd x" and "y \ CH (fp_cop_F ds)" by clarsimp with assms \x \ X''\ show ?thesis using CH_range fp_cop_F_closed_inv' fp_cop_F_preferred unfolding above_def by blast next case False with assms \x \ X''\ show ?thesis by (meson Cd_range' IntD2 fp_cop_F_all mem_CD_on_Cd rev_subsetD) qed qed text (in Contracts) \ The @{const "irc"} constraint on the auctioneer's preferences is needed for @{const "stable_no_blocking"} and their part of @{const "individually_rational"}. \ end context ContractsWithIRC begin lemma cop_stable_no_blocking_on: shows "stable_no_blocking_on ds (cop ds)" proof(rule stable_no_blocking_onI) fix h X'' assume C: "X'' = Ch h (CH (fp_cop_F ds) \ X'')" assume NE: "X'' \ Ch h (CH (fp_cop_F ds))" assume CD: "X'' \ CD_on ds (CH (fp_cop_F ds) \ X'')" from CD have "X'' \ fp_cop_F ds" by (rule X''_closed) then have X: "CH (fp_cop_F ds) \ X'' \ fp_cop_F ds" using CH_range by simp from C NE Ch_CH_irc_idem[of h] show False using consistency_onD[OF Ch_consistency _ X] CH_domain Ch_domain by blast qed theorem Theorem_16: assumes h: "(UNIV::'c set) = {h}" shows "stable_on ds (cop ds)" (is "stable_on ds ?fp") proof(rule stable_onI) show "individually_rational_on ds ?fp" proof(rule individually_rational_onI) from h have "allocation ?fp" by (simp add: Ch_singular CH_Ch_singular) then show "CD_on ds ?fp = ?fp" by (rule CD_on_closed) (blast dest: CH_range' fp_cop_F_range_inv') show "CH (CH (fp_cop_F ds)) = CH (fp_cop_F ds)" by (simp add: CH_irc_idem) qed show "stable_no_blocking_on ds ?fp" by (rule cop_stable_no_blocking_on) qed end subsection\ Concluding remarks \ text\ From \citet{HatfieldMilgrom:2005}, we have not shown Theorems~2, 7, 13 and~14, all of which are intended to position their results against prior work in this space. We delay establishing their strategic results (Theorems~10, 11 and~12) to \S\ref{sec:strategic}, after we have developed more useful invariants for the COP. By assuming \isa{irc}, \citet{AygunSonmez:2012-WP2} are essentially trading on Plott's path independence condition (\S\ref{sec:cf-path-independence}), as observed by \citet{ChambersYenmez:2013}. The latter show that these results generalize naturally to many-to-many matches, where doctors also use path-independent choice functions; see also \citet{Fleiner:2003}. For many applications, however, @{const "substitutes"} proves to be too strong a condition. The COP of \S\ref{sec:contracts-cop} provides a way forward, as we discuss in the next section. \ (*<*) end (*>*)