# HG changeset patch # User Xingyuan Zhang # Date 1410574034 -28800 # Node ID a5f5b9336007270ab826d6d41228d967d317b8d2 # Parent 77daf1b85cf09ec77b3b0a1eab4b6bb8eaab5efb thys2 added diff -r 77daf1b85cf0 -r a5f5b9336007 ROOT --- a/ROOT Fri Sep 12 00:47:15 2014 +0800 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,26 +0,0 @@ -session "Hoare_gen" in "thys" = "HOL" + - theories - Hoare_gen - -session "Hoare_tm" in "thys" = "Hoare_gen" + - theories - Hoare_tm - Hoare_tm2 - -session "Hoare_abc" in "thys" = "Hoare_tm" + - theories - Hoare_abc - Hoare_abc2 - -session "TM" in "thys" = Hoare_abc + - theories - My_block - LetElim - StateMonad - TM_Assemble - -session "Paper" in "paper" = TM + - options [document = pdf, document_output = "..", document_variants = "journal"] - theories - Paper - files "document/root.tex" \ No newline at end of file diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/Map_Extra.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/Map_Extra.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,573 @@ +(* Author: Rafal Kolanski, 2008 + Maintainers: Gerwin Klein + Rafal Kolanski +*) + +header {* More properties of maps plus map disjuction. *} + +theory Map_Extra + imports Main +begin + +text {* + A note on naming: + Anything not involving heap disjuction can potentially be incorporated + directly into Map.thy, thus uses @{text "m"} for map variable names. + Anything involving heap disjunction is not really mergeable with Map, is + destined for use in separation logic, and hence uses @{text "h"} +*} + +section {* Things that could go into Option Type *} + +text {* Misc option lemmas *} + +lemma None_not_eq: "(None \ x) = (\y. x = Some y)" by (cases x) auto + +lemma None_com: "(None = x) = (x = None)" by fast + +lemma Some_com: "(Some y = x) = (x = Some y)" by fast + + +section {* Things that go into Map.thy *} + +text {* Map intersection: set of all keys for which the maps agree. *} + +definition + map_inter :: "('a \ 'b) \ ('a \ 'b) \ 'a set" (infixl "\\<^isub>m" 70) where + "m\<^isub>1 \\<^isub>m m\<^isub>2 \ {x \ dom m\<^isub>1. m\<^isub>1 x = m\<^isub>2 x}" + +text {* Map restriction via domain subtraction *} + +definition + sub_restrict_map :: "('a \ 'b) => 'a set => ('a \ 'b)" (infixl "`-" 110) + where + "m `- S \ (\x. if x \ S then None else m x)" + + +subsection {* Properties of maps not related to restriction *} + +lemma empty_forall_equiv: "(m = empty) = (\x. m x = None)" + by (fastforce intro!: ext) + +lemma map_le_empty2 [simp]: + "(m \\<^sub>m empty) = (m = empty)" + by (auto simp: map_le_def intro: ext) + +lemma dom_iff: + "(\y. m x = Some y) = (x \ dom m)" + by auto + +lemma non_dom_eval: + "x \ dom m \ m x = None" + by auto + +lemma non_dom_eval_eq: + "x \ dom m = (m x = None)" + by auto + +lemma map_add_same_left_eq: + "m\<^isub>1 = m\<^isub>1' \ (m\<^isub>0 ++ m\<^isub>1 = m\<^isub>0 ++ m\<^isub>1')" + by simp + +lemma map_add_left_cancelI [intro!]: + "m\<^isub>1 = m\<^isub>1' \ m\<^isub>0 ++ m\<^isub>1 = m\<^isub>0 ++ m\<^isub>1'" + by simp + +lemma dom_empty_is_empty: + "(dom m = {}) = (m = empty)" +proof (rule iffI) + assume a: "dom m = {}" + { assume "m \ empty" + hence "dom m \ {}" + by - (subst (asm) empty_forall_equiv, simp add: dom_def) + hence False using a by blast + } + thus "m = empty" by blast +next + assume a: "m = empty" + thus "dom m = {}" by simp +qed + +lemma map_add_dom_eq: + "dom m = dom m' \ m ++ m' = m'" + by (rule ext) (auto simp: map_add_def split: option.splits) + +lemma map_add_right_dom_eq: + "\ m\<^isub>0 ++ m\<^isub>1 = m\<^isub>0' ++ m\<^isub>1'; dom m\<^isub>1 = dom m\<^isub>1' \ \ m\<^isub>1 = m\<^isub>1'" + unfolding map_add_def + by (rule ext, rule ccontr, + drule_tac x=x in fun_cong, clarsimp split: option.splits, + drule sym, drule sym, force+) + +lemma map_le_same_dom_eq: + "\ m\<^isub>0 \\<^sub>m m\<^isub>1 ; dom m\<^isub>0 = dom m\<^isub>1 \ \ m\<^isub>0 = m\<^isub>1" + by (auto intro!: ext simp: map_le_def elim!: ballE) + + +subsection {* Properties of map restriction *} + +lemma restrict_map_cancel: + "(m |` S = m |` T) = (dom m \ S = dom m \ T)" + by (fastforce intro: ext dest: fun_cong + simp: restrict_map_def None_not_eq + split: split_if_asm) + +lemma map_add_restricted_self [simp]: + "m ++ m |` S = m" + by (auto intro: ext simp: restrict_map_def map_add_def split: option.splits) + +lemma map_add_restrict_dom_right [simp]: + "(m ++ m') |` dom m' = m'" + by (rule ext, auto simp: restrict_map_def map_add_def split: option.splits) + +lemma restrict_map_UNIV [simp]: + "m |` UNIV = m" + by (simp add: restrict_map_def) + +lemma restrict_map_dom: + "S = dom m \ m |` S = m" + by (auto intro!: ext simp: restrict_map_def None_not_eq) + +lemma restrict_map_subdom: + "dom m \ S \ m |` S = m" + by (fastforce simp: restrict_map_def None_com intro: ext) + +lemma map_add_restrict: + "(m\<^isub>0 ++ m\<^isub>1) |` S = ((m\<^isub>0 |` S) ++ (m\<^isub>1 |` S))" + by (force simp: map_add_def restrict_map_def intro: ext) + +lemma map_le_restrict: + "m \\<^sub>m m' \ m = m' |` dom m" + by (force simp: map_le_def restrict_map_def None_com intro: ext) + +lemma restrict_map_le: + "m |` S \\<^sub>m m" + by (auto simp: map_le_def) + +lemma restrict_map_remerge: + "\ S \ T = {} \ \ m |` S ++ m |` T = m |` (S \ T)" + by (rule ext, clarsimp simp: restrict_map_def map_add_def + split: option.splits) + +lemma restrict_map_empty: + "dom m \ S = {} \ m |` S = empty" + by (fastforce simp: restrict_map_def intro: ext) + +lemma map_add_restrict_comp_right [simp]: + "(m |` S ++ m |` (UNIV - S)) = m" + by (force simp: map_add_def restrict_map_def split: option.splits intro: ext) + +lemma map_add_restrict_comp_right_dom [simp]: + "(m |` S ++ m |` (dom m - S)) = m" + by (auto simp: map_add_def restrict_map_def split: option.splits intro!: ext) + +lemma map_add_restrict_comp_left [simp]: + "(m |` (UNIV - S) ++ m |` S) = m" + by (subst map_add_comm, auto) + +lemma restrict_self_UNIV: + "m |` (dom m - S) = m |` (UNIV - S)" + by (auto intro!: ext simp: restrict_map_def) + +lemma map_add_restrict_nonmember_right: + "x \ dom m' \ (m ++ m') |` {x} = m |` {x}" + by (rule ext, auto simp: restrict_map_def map_add_def split: option.splits) + +lemma map_add_restrict_nonmember_left: + "x \ dom m \ (m ++ m') |` {x} = m' |` {x}" + by (rule ext, auto simp: restrict_map_def map_add_def split: option.splits) + +lemma map_add_restrict_right: + "x \ dom m' \ (m ++ m') |` x = m' |` x" + by (rule ext, auto simp: restrict_map_def map_add_def split: option.splits) + +lemma restrict_map_compose: + "\ S \ T = dom m ; S \ T = {} \ \ m |` S ++ m |` T = m" + by (fastforce intro: ext simp: map_add_def restrict_map_def) + +lemma map_le_dom_subset_restrict: + "\ m' \\<^sub>m m; dom m' \ S \ \ m' \\<^sub>m (m |` S)" + by (force simp: restrict_map_def map_le_def) + +lemma map_le_dom_restrict_sub_add: + "m' \\<^sub>m m \ m |` (dom m - dom m') ++ m' = m" + by (auto simp: None_com map_add_def restrict_map_def map_le_def + split: option.splits + intro!: ext) + (force simp: Some_com)+ + +lemma subset_map_restrict_sub_add: + "T \ S \ m |` (S - T) ++ m |` T = m |` S" + by (auto simp: restrict_map_def map_add_def intro!: ext split: option.splits) + +lemma restrict_map_sub_union: + "m |` (dom m - (S \ T)) = (m |` (dom m - T)) |` (dom m - S)" + by (auto intro!: ext simp: restrict_map_def) + +lemma prod_restrict_map_add: + "\ S \ T = U; S \ T = {} \ \ m |` (X \ S) ++ m |` (X \ T) = m |` (X \ U)" + by (auto simp: map_add_def restrict_map_def intro!: ext split: option.splits) + + +section {* Things that should not go into Map.thy (separation logic) *} + +subsection {* Definitions *} + +text {* Map disjuction *} + +definition + map_disj :: "('a \ 'b) \ ('a \ 'b) \ bool" (infix "\" 51) where + "h\<^isub>0 \ h\<^isub>1 \ dom h\<^isub>0 \ dom h\<^isub>1 = {}" + +declare None_not_eq [simp] + + +subsection {* Properties of @{term "sub_restrict_map"} *} + +lemma restrict_map_sub_disj: "h |` S \ h `- S" + by (fastforce simp: sub_restrict_map_def restrict_map_def map_disj_def + split: option.splits split_if_asm) + +lemma restrict_map_sub_add: "h |` S ++ h `- S = h" + by (fastforce simp: sub_restrict_map_def restrict_map_def map_add_def + split: option.splits split_if + intro: ext) + + +subsection {* Properties of map disjunction *} + +lemma map_disj_empty_right [simp]: + "h \ empty" + by (simp add: map_disj_def) + +lemma map_disj_empty_left [simp]: + "empty \ h" + by (simp add: map_disj_def) + +lemma map_disj_com: + "h\<^isub>0 \ h\<^isub>1 = h\<^isub>1 \ h\<^isub>0" + by (simp add: map_disj_def, fast) + +lemma map_disjD: + "h\<^isub>0 \ h\<^isub>1 \ dom h\<^isub>0 \ dom h\<^isub>1 = {}" + by (simp add: map_disj_def) + +lemma map_disjI: + "dom h\<^isub>0 \ dom h\<^isub>1 = {} \ h\<^isub>0 \ h\<^isub>1" + by (simp add: map_disj_def) + + +subsection {* Map associativity-commutativity based on map disjuction *} + +lemma map_add_com: + "h\<^isub>0 \ h\<^isub>1 \ h\<^isub>0 ++ h\<^isub>1 = h\<^isub>1 ++ h\<^isub>0" + by (drule map_disjD, rule map_add_comm, force) + +lemma map_add_left_commute: + "h\<^isub>0 \ h\<^isub>1 \ h\<^isub>0 ++ (h\<^isub>1 ++ h\<^isub>2) = h\<^isub>1 ++ (h\<^isub>0 ++ h\<^isub>2)" + by (simp add: map_add_com map_disj_com map_add_assoc) + +lemma map_add_disj: + "h\<^isub>0 \ (h\<^isub>1 ++ h\<^isub>2) = (h\<^isub>0 \ h\<^isub>1 \ h\<^isub>0 \ h\<^isub>2)" + by (simp add: map_disj_def, fast) + +lemma map_add_disj': + "(h\<^isub>1 ++ h\<^isub>2) \ h\<^isub>0 = (h\<^isub>1 \ h\<^isub>0 \ h\<^isub>2 \ h\<^isub>0)" + by (simp add: map_disj_def, fast) + +text {* + We redefine @{term "map_add"} associativity to bind to the right, which + seems to be the more common case. + Note that when a theory includes Map again, @{text "map_add_assoc"} will + return to the simpset and will cause infinite loops if its symmetric + counterpart is added (e.g. via @{text "map_add_ac"}) + *} + +declare map_add_assoc [simp del] + +text {* + Since the associativity-commutativity of @{term "map_add"} relies on + map disjunction, we include some basic rules into the ac set. + *} + +lemmas map_add_ac = + map_add_assoc[symmetric] map_add_com map_disj_com + map_add_left_commute map_add_disj map_add_disj' + + +subsection {* Basic properties *} + +lemma map_disj_None_right: + "\ h\<^isub>0 \ h\<^isub>1 ; x \ dom h\<^isub>0 \ \ h\<^isub>1 x = None" + by (auto simp: map_disj_def dom_def) + +lemma map_disj_None_left: + "\ h\<^isub>0 \ h\<^isub>1 ; x \ dom h\<^isub>1 \ \ h\<^isub>0 x = None" + by (auto simp: map_disj_def dom_def) + +lemma map_disj_None_left': + "\ h\<^isub>0 x = Some y ; h\<^isub>1 \ h\<^isub>0 \ \ h\<^isub>1 x = None " + by (auto simp: map_disj_def) + +lemma map_disj_None_right': + "\ h\<^isub>1 x = Some y ; h\<^isub>1 \ h\<^isub>0 \ \ h\<^isub>0 x = None " + by (auto simp: map_disj_def) + +lemma map_disj_common: + "\ h\<^isub>0 \ h\<^isub>1 ; h\<^isub>0 p = Some v ; h\<^isub>1 p = Some v' \ \ False" + by (frule (1) map_disj_None_left', simp) + +lemma map_disj_eq_dom_left: + "\ h\<^isub>0 \ h\<^isub>1 ; dom h\<^isub>0' = dom h\<^isub>0 \ \ h\<^isub>0' \ h\<^isub>1" + by (auto simp: map_disj_def) + + +subsection {* Map disjunction and addition *} + +lemma map_add_eval_left: + "\ x \ dom h ; h \ h' \ \ (h ++ h') x = h x" + by (auto dest!: map_disj_None_right simp: map_add_def cong: option.case_cong) + +lemma map_add_eval_right: + "\ x \ dom h' ; h \ h' \ \ (h ++ h') x = h' x" + by (auto elim!: map_disjD simp: map_add_comm map_add_eval_left map_disj_com) + +lemma map_add_eval_left': + "\ x \ dom h' ; h \ h' \ \ (h ++ h') x = h x" + by (clarsimp simp: map_disj_def map_add_def split: option.splits) + +lemma map_add_eval_right': + "\ x \ dom h ; h \ h' \ \ (h ++ h') x = h' x" + by (clarsimp simp: map_disj_def map_add_def split: option.splits) + +lemma map_add_left_dom_eq: + assumes eq: "h\<^isub>0 ++ h\<^isub>1 = h\<^isub>0' ++ h\<^isub>1'" + assumes etc: "h\<^isub>0 \ h\<^isub>1" "h\<^isub>0' \ h\<^isub>1'" "dom h\<^isub>0 = dom h\<^isub>0'" + shows "h\<^isub>0 = h\<^isub>0'" +proof - + from eq have "h\<^isub>1 ++ h\<^isub>0 = h\<^isub>1' ++ h\<^isub>0'" using etc by (simp add: map_add_ac) + thus ?thesis using etc + by (fastforce elim!: map_add_right_dom_eq simp: map_add_ac) +qed + +lemma map_add_left_eq: + assumes eq: "h\<^isub>0 ++ h = h\<^isub>1 ++ h" + assumes disj: "h\<^isub>0 \ h" "h\<^isub>1 \ h" + shows "h\<^isub>0 = h\<^isub>1" +proof (rule ext) + fix x + from eq have eq': "(h\<^isub>0 ++ h) x = (h\<^isub>1 ++ h) x" by (auto intro!: ext) + { assume "x \ dom h" + hence "h\<^isub>0 x = h\<^isub>1 x" using disj by (simp add: map_disj_None_left) + } moreover { + assume "x \ dom h" + hence "h\<^isub>0 x = h\<^isub>1 x" using disj eq' by (simp add: map_add_eval_left') + } + ultimately show "h\<^isub>0 x = h\<^isub>1 x" by cases +qed + +lemma map_add_right_eq: + "\h ++ h\<^isub>0 = h ++ h\<^isub>1; h\<^isub>0 \ h; h\<^isub>1 \ h\ \ h\<^isub>0 = h\<^isub>1" + by (rule_tac h=h in map_add_left_eq, auto simp: map_add_ac) + +lemma map_disj_add_eq_dom_right_eq: + assumes merge: "h\<^isub>0 ++ h\<^isub>1 = h\<^isub>0' ++ h\<^isub>1'" and d: "dom h\<^isub>0 = dom h\<^isub>0'" and + ab_disj: "h\<^isub>0 \ h\<^isub>1" and cd_disj: "h\<^isub>0' \ h\<^isub>1'" + shows "h\<^isub>1 = h\<^isub>1'" +proof (rule ext) + fix x + from merge have merge_x: "(h\<^isub>0 ++ h\<^isub>1) x = (h\<^isub>0' ++ h\<^isub>1') x" by simp + with d ab_disj cd_disj show "h\<^isub>1 x = h\<^isub>1' x" + by - (case_tac "h\<^isub>1 x", case_tac "h\<^isub>1' x", simp, fastforce simp: map_disj_def, + case_tac "h\<^isub>1' x", clarsimp, simp add: Some_com, + force simp: map_disj_def, simp) +qed + +lemma map_disj_add_eq_dom_left_eq: + assumes add: "h\<^isub>0 ++ h\<^isub>1 = h\<^isub>0' ++ h\<^isub>1'" and + dom: "dom h\<^isub>1 = dom h\<^isub>1'" and + disj: "h\<^isub>0 \ h\<^isub>1" "h\<^isub>0' \ h\<^isub>1'" + shows "h\<^isub>0 = h\<^isub>0'" +proof - + have "h\<^isub>1 ++ h\<^isub>0 = h\<^isub>1' ++ h\<^isub>0'" using add disj by (simp add: map_add_ac) + thus ?thesis using dom disj + by - (rule map_disj_add_eq_dom_right_eq, auto simp: map_disj_com) +qed + +lemma map_add_left_cancel: + assumes disj: "h\<^isub>0 \ h\<^isub>1" "h\<^isub>0 \ h\<^isub>1'" + shows "(h\<^isub>0 ++ h\<^isub>1 = h\<^isub>0 ++ h\<^isub>1') = (h\<^isub>1 = h\<^isub>1')" +proof (rule iffI, rule ext) + fix x + assume "(h\<^isub>0 ++ h\<^isub>1) = (h\<^isub>0 ++ h\<^isub>1')" + hence "(h\<^isub>0 ++ h\<^isub>1) x = (h\<^isub>0 ++ h\<^isub>1') x" by (auto intro!: ext) + hence "h\<^isub>1 x = h\<^isub>1' x" using disj + by - (cases "x \ dom h\<^isub>0", + simp_all add: map_disj_None_right map_add_eval_right') + thus "h\<^isub>1 x = h\<^isub>1' x" by (auto intro!: ext) +qed auto + +lemma map_add_lr_disj: + "\ h\<^isub>0 ++ h\<^isub>1 = h\<^isub>0' ++ h\<^isub>1'; h\<^isub>1 \ h\<^isub>1' \ \ dom h\<^isub>1 \ dom h\<^isub>0'" + by (clarsimp simp: map_disj_def map_add_def, drule_tac x=x in fun_cong) + (auto split: option.splits) + + +subsection {* Map disjunction and map updates *} + +lemma map_disj_update_left [simp]: + "p \ dom h\<^isub>1 \ h\<^isub>0 \ h\<^isub>1(p \ v) = h\<^isub>0 \ h\<^isub>1" + by (clarsimp simp add: map_disj_def, blast) + +lemma map_disj_update_right [simp]: + "p \ dom h\<^isub>1 \ h\<^isub>1(p \ v) \ h\<^isub>0 = h\<^isub>1 \ h\<^isub>0" + by (simp add: map_disj_com) + +lemma map_add_update_left: + "\ h\<^isub>0 \ h\<^isub>1 ; p \ dom h\<^isub>0 \ \ (h\<^isub>0 ++ h\<^isub>1)(p \ v) = (h\<^isub>0(p \ v) ++ h\<^isub>1)" + by (drule (1) map_disj_None_right) + (auto intro: ext simp: map_add_def cong: option.case_cong) + +lemma map_add_update_right: + "\ h\<^isub>0 \ h\<^isub>1 ; p \ dom h\<^isub>1 \ \ (h\<^isub>0 ++ h\<^isub>1)(p \ v) = (h\<^isub>0 ++ h\<^isub>1 (p \ v))" + by (drule (1) map_disj_None_left) + (auto intro: ext simp: map_add_def cong: option.case_cong) + +lemma map_add3_update: + "\ h\<^isub>0 \ h\<^isub>1 ; h\<^isub>1 \ h\<^isub>2 ; h\<^isub>0 \ h\<^isub>2 ; p \ dom h\<^isub>0 \ + \ (h\<^isub>0 ++ h\<^isub>1 ++ h\<^isub>2)(p \ v) = h\<^isub>0(p \ v) ++ h\<^isub>1 ++ h\<^isub>2" + by (auto simp: map_add_update_left[symmetric] map_add_ac) + + +subsection {* Map disjunction and @{term "map_le"} *} + +lemma map_le_override [simp]: + "\ h \ h' \ \ h \\<^sub>m h ++ h'" + by (auto simp: map_le_def map_add_def map_disj_def split: option.splits) + +lemma map_leI_left: + "\ h = h\<^isub>0 ++ h\<^isub>1 ; h\<^isub>0 \ h\<^isub>1 \ \ h\<^isub>0 \\<^sub>m h" by auto + +lemma map_leI_right: + "\ h = h\<^isub>0 ++ h\<^isub>1 ; h\<^isub>0 \ h\<^isub>1 \ \ h\<^isub>1 \\<^sub>m h" by auto + +lemma map_disj_map_le: + "\ h\<^isub>0' \\<^sub>m h\<^isub>0; h\<^isub>0 \ h\<^isub>1 \ \ h\<^isub>0' \ h\<^isub>1" + by (force simp: map_disj_def map_le_def) + +lemma map_le_on_disj_left: + "\ h' \\<^sub>m h ; h\<^isub>0 \ h\<^isub>1 ; h' = h\<^isub>0 ++ h\<^isub>1 \ \ h\<^isub>0 \\<^sub>m h" + unfolding map_le_def + by (rule ballI, erule_tac x=a in ballE, auto simp: map_add_eval_left)+ + +lemma map_le_on_disj_right: + "\ h' \\<^sub>m h ; h\<^isub>0 \ h\<^isub>1 ; h' = h\<^isub>1 ++ h\<^isub>0 \ \ h\<^isub>0 \\<^sub>m h" + by (auto simp: map_le_on_disj_left map_add_ac) + +lemma map_le_add_cancel: + "\ h\<^isub>0 \ h\<^isub>1 ; h\<^isub>0' \\<^sub>m h\<^isub>0 \ \ h\<^isub>0' ++ h\<^isub>1 \\<^sub>m h\<^isub>0 ++ h\<^isub>1" + by (auto simp: map_le_def map_add_def map_disj_def split: option.splits) + +lemma map_le_override_bothD: + assumes subm: "h\<^isub>0' ++ h\<^isub>1 \\<^sub>m h\<^isub>0 ++ h\<^isub>1" + assumes disj': "h\<^isub>0' \ h\<^isub>1" + assumes disj: "h\<^isub>0 \ h\<^isub>1" + shows "h\<^isub>0' \\<^sub>m h\<^isub>0" +unfolding map_le_def +proof (rule ballI) + fix a + assume a: "a \ dom h\<^isub>0'" + hence sumeq: "(h\<^isub>0' ++ h\<^isub>1) a = (h\<^isub>0 ++ h\<^isub>1) a" + using subm unfolding map_le_def by auto + from a have "a \ dom h\<^isub>1" using disj' by (auto dest!: map_disj_None_right) + thus "h\<^isub>0' a = h\<^isub>0 a" using a sumeq disj disj' + by (simp add: map_add_eval_left map_add_eval_left') +qed + +lemma map_le_conv: + "(h\<^isub>0' \\<^sub>m h\<^isub>0 \ h\<^isub>0' \ h\<^isub>0) = (\h\<^isub>1. h\<^isub>0 = h\<^isub>0' ++ h\<^isub>1 \ h\<^isub>0' \ h\<^isub>1 \ h\<^isub>0' \ h\<^isub>0)" + unfolding map_le_def map_disj_def map_add_def + by (rule iffI, + clarsimp intro!: exI[where x="\x. if x \ dom h\<^isub>0' then h\<^isub>0 x else None"]) + (fastforce intro: ext intro: split: option.splits split_if_asm)+ + +lemma map_le_conv2: + "h\<^isub>0' \\<^sub>m h\<^isub>0 = (\h\<^isub>1. h\<^isub>0 = h\<^isub>0' ++ h\<^isub>1 \ h\<^isub>0' \ h\<^isub>1)" + by (case_tac "h\<^isub>0'=h\<^isub>0", insert map_le_conv, auto intro: exI[where x=empty]) + + +subsection {* Map disjunction and restriction *} + +lemma map_disj_comp [simp]: + "h\<^isub>0 \ h\<^isub>1 |` (UNIV - dom h\<^isub>0)" + by (force simp: map_disj_def) + +lemma restrict_map_disj: + "S \ T = {} \ h |` S \ h |` T" + by (auto simp: map_disj_def restrict_map_def dom_def) + +lemma map_disj_restrict_dom [simp]: + "h\<^isub>0 \ h\<^isub>1 |` (dom h\<^isub>1 - dom h\<^isub>0)" + by (force simp: map_disj_def) + +lemma restrict_map_disj_dom_empty: + "h \ h' \ h |` dom h' = empty" + by (fastforce simp: map_disj_def restrict_map_def intro: ext) + +lemma restrict_map_univ_disj_eq: + "h \ h' \ h |` (UNIV - dom h') = h" + by (rule ext, auto simp: map_disj_def restrict_map_def) + +lemma restrict_map_disj_dom: + "h\<^isub>0 \ h\<^isub>1 \ h |` dom h\<^isub>0 \ h |` dom h\<^isub>1" + by (auto simp: map_disj_def restrict_map_def dom_def) + +lemma map_add_restrict_dom_left: + "h \ h' \ (h ++ h') |` dom h = h" + by (rule ext, auto simp: restrict_map_def map_add_def dom_def map_disj_def + split: option.splits) + +lemma map_add_restrict_dom_left': + "h \ h' \ S = dom h \ (h ++ h') |` S = h" + by (rule ext, auto simp: restrict_map_def map_add_def dom_def map_disj_def + split: option.splits) + +lemma restrict_map_disj_left: + "h\<^isub>0 \ h\<^isub>1 \ h\<^isub>0 |` S \ h\<^isub>1" + by (auto simp: map_disj_def) + +lemma restrict_map_disj_right: + "h\<^isub>0 \ h\<^isub>1 \ h\<^isub>0 \ h\<^isub>1 |` S" + by (auto simp: map_disj_def) + +lemmas restrict_map_disj_both = restrict_map_disj_right restrict_map_disj_left + +lemma map_dom_disj_restrict_right: + "h\<^isub>0 \ h\<^isub>1 \ (h\<^isub>0 ++ h\<^isub>0') |` dom h\<^isub>1 = h\<^isub>0' |` dom h\<^isub>1" + by (simp add: map_add_restrict restrict_map_empty map_disj_def) + +lemma restrict_map_on_disj: + "h\<^isub>0' \ h\<^isub>1 \ h\<^isub>0 |` dom h\<^isub>0' \ h\<^isub>1" + unfolding map_disj_def by auto + +lemma restrict_map_on_disj': + "h\<^isub>0 \ h\<^isub>1 \ h\<^isub>0 \ h\<^isub>1 |` S" + by (auto simp: map_disj_def map_add_def) + +lemma map_le_sub_dom: + "\ h\<^isub>0 ++ h\<^isub>1 \\<^sub>m h ; h\<^isub>0 \ h\<^isub>1 \ \ h\<^isub>0 \\<^sub>m h |` (dom h - dom h\<^isub>1)" + by (rule map_le_override_bothD, subst map_le_dom_restrict_sub_add) + (auto elim: map_add_le_mapE simp: map_add_ac) + +lemma map_submap_break: + "\ h \\<^sub>m h' \ \ h' = (h' |` (UNIV - dom h)) ++ h" + by (fastforce intro!: ext split: option.splits + simp: map_le_restrict restrict_map_def map_le_def map_add_def + dom_def) + +lemma map_add_disj_restrict_both: + "\ h\<^isub>0 \ h\<^isub>1; S \ S' = {}; T \ T' = {} \ + \ (h\<^isub>0 |` S) ++ (h\<^isub>1 |` T) \ (h\<^isub>0 |` S') ++ (h\<^isub>1 |` T')" + by (auto simp: map_add_ac intro!: restrict_map_disj_both restrict_map_disj) + +end diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/README --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/README Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,11 @@ +This directory contains a generic type class implementation of separation +algebra for Isabelle/HOL as well as lemmas and generic tactics which can be +used directly for any instantiation of the type class. + +The ex directory contains example instantiations that include structures such +as a heap or virtual memory. + +The abstract separation algebra is based upon "Abstract Separation Logic" by +Calcagno et al. These theories are also the basis of "Mechanised Separation +Algebra" by Klein et al. + diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/ROOT --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/ROOT Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,11 @@ +session Separation_Algebra (AFP) = "HOL-Word" + + options [timeout = 3600] + theories + "ex/Simple_Separation_Example" + "ex/Sep_Tactics_Test" + "ex/VM_Example" + Sep_Eq + "ex/capDL/Separation_D" + files + "document/root.bib" + "document/root.tex" diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/Sep_Eq.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/Sep_Eq.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,108 @@ +(* Author: Gerwin Klein, 2012 + Maintainers: Gerwin Klein + Rafal Kolanski +*) + +header "Equivalence between Separation Algebra Formulations" + +theory Sep_Eq +imports Separation_Algebra Separation_Algebra_Alt +begin + +text {* + In this theory we show that our total formulation of separation algebra is + equivalent in strength to Calcagno et al's original partial one. + + This theory is not intended to be included in own developments. +*} + +no_notation map_add (infixl "++" 100) + +section "Total implies Partial" + +definition add2 :: "'a::sep_algebra => 'a => 'a option" where + "add2 x y \ if x ## y then Some (x + y) else None" + +lemma add2_zero: "add2 x 0 = Some x" + by (simp add: add2_def) + +lemma add2_comm: "add2 x y = add2 y x" + by (auto simp: add2_def sep_add_commute sep_disj_commute) + +lemma add2_assoc: + "lift2 add2 a (lift2 add2 b c) = lift2 add2 (lift2 add2 a b) c" + by (auto simp: add2_def lift2_def sep_add_assoc + dest: sep_disj_addD sep_disj_addI3 + sep_add_disjD sep_disj_addI2 sep_disj_commuteI + split: option.splits) + +interpretation total_partial: sep_algebra_alt 0 add2 + by (unfold_locales) (auto intro: add2_zero add2_comm add2_assoc) + + +section "Partial implies Total" + +definition + sep_add' :: "'a \ 'a \ 'a :: sep_algebra_alt" where + "sep_add' x y \ if disjoint x y then the (add x y) else undefined" + +lemma sep_disj_zero': + "disjoint x 0" + by simp + +lemma sep_disj_commuteI': + "disjoint x y \ disjoint y x" + by (clarsimp simp: disjoint_def add_comm) + +lemma sep_add_zero': + "sep_add' x 0 = x" + by (simp add: sep_add'_def) + +lemma sep_add_commute': + "disjoint x y \ sep_add' x y = sep_add' y x" + by (clarsimp simp: sep_add'_def disjoint_def add_comm) + +lemma sep_add_assoc': + "\ disjoint x y; disjoint y z; disjoint x z \ \ + sep_add' (sep_add' x y) z = sep_add' x (sep_add' y z)" + using add_assoc [of "Some x" "Some y" "Some z"] + by (clarsimp simp: disjoint_def sep_add'_def lift2_def + split: option.splits) + +lemma sep_disj_addD1': + "disjoint x (sep_add' y z) \ disjoint y z \ disjoint x y" +proof (clarsimp simp: disjoint_def sep_add'_def) + fix a assume a: "y \ z = Some a" + fix b assume b: "x \ a = Some b" + with a have "Some x ++ (Some y ++ Some z) = Some b" by (simp add: lift2_def) + hence "(Some x ++ Some y) ++ Some z = Some b" by (simp add: add_assoc) + thus "\b. x \ y = Some b" by (simp add: lift2_def split: option.splits) +qed + +lemma sep_disj_addI1': + "disjoint x (sep_add' y z) \ disjoint y z \ disjoint (sep_add' x y) z" + apply (clarsimp simp: disjoint_def sep_add'_def) + apply (rule conjI) + apply clarsimp + apply (frule lift_to_add2, assumption) + apply (simp add: add_assoc) + apply (clarsimp simp: lift2_def add_comm) + apply clarsimp + apply (frule lift_to_add2, assumption) + apply (simp add: add_assoc) + apply (clarsimp simp: lift2_def) + done + +interpretation partial_total: sep_algebra sep_add' 0 disjoint + apply (unfold_locales) + apply (rule sep_disj_zero') + apply (erule sep_disj_commuteI') + apply (rule sep_add_zero') + apply (erule sep_add_commute') + apply (erule (2) sep_add_assoc') + apply (erule (1) sep_disj_addD1') + apply (erule (1) sep_disj_addI1') + done + +end + diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/Sep_Heap_Instance.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/Sep_Heap_Instance.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,84 @@ +(* Author: Gerwin Klein, 2012 + Maintainers: Gerwin Klein + Rafal Kolanski +*) + +header "Standard Heaps as an Instance of Separation Algebra" + +theory Sep_Heap_Instance +imports Separation_Algebra +begin + +text {* + Example instantiation of a the separation algebra to a map, i.e.\ a + function from any type to @{typ "'a option"}. +*} + +class opt = + fixes none :: 'a +begin + definition "domain f \ {x. f x \ none}" +end + +instantiation option :: (type) opt +begin + definition none_def [simp]: "none \ None" + instance .. +end + +instantiation "fun" :: (type, opt) zero +begin + definition zero_fun_def: "0 \ \s. none" + instance .. +end + +instantiation "fun" :: (type, opt) sep_algebra +begin + +definition + plus_fun_def: "m1 + m2 \ \x. if m2 x = none then m1 x else m2 x" + +definition + sep_disj_fun_def: "sep_disj m1 m2 \ domain m1 \ domain m2 = {}" + +instance + apply default + apply (simp add: sep_disj_fun_def domain_def zero_fun_def) + apply (fastforce simp: sep_disj_fun_def) + apply (simp add: plus_fun_def zero_fun_def) + apply (simp add: plus_fun_def sep_disj_fun_def domain_def) + apply (rule ext) + apply fastforce + apply (rule ext) + apply (simp add: plus_fun_def) + apply (simp add: sep_disj_fun_def domain_def plus_fun_def) + apply fastforce + apply (simp add: sep_disj_fun_def domain_def plus_fun_def) + apply fastforce + done + +end + +text {* + For the actual option type @{const domain} and @{text "+"} are + just @{const dom} and @{text "++"}: +*} + +lemma domain_conv: "domain = dom" + by (rule ext) (simp add: domain_def dom_def) + +lemma plus_fun_conv: "a + b = a ++ b" + by (auto simp: plus_fun_def map_add_def split: option.splits) + +lemmas map_convs = domain_conv plus_fun_conv + +text {* + Any map can now act as a separation heap without further work: +*} +lemma + fixes h :: "(nat => nat) => 'foo option" + shows "(P ** Q ** H) h = (Q ** H ** P) h" + by (simp add: sep_conj_ac) + +end + diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/Sep_Tactics.thy --- a/Separation_Algebra/Sep_Tactics.thy Fri Sep 12 00:47:15 2014 +0800 +++ b/Separation_Algebra/Sep_Tactics.thy Sat Sep 13 10:07:14 2014 +0800 @@ -6,50 +6,78 @@ header "Separation Logic Tactics" theory Sep_Tactics -imports Separation_Algebra +imports Separation_Algebra +uses "sep_tactics.ML" begin -ML_file "sep_tactics.ML" text {* A number of proof methods to assist with reasoning about separation logic. *} section {* Selection (move-to-front) tactics *} +ML {* + fun sep_select_method n ctxt = + Method.SIMPLE_METHOD' (sep_select_tac ctxt n); + fun sep_select_asm_method n ctxt = + Method.SIMPLE_METHOD' (sep_select_asm_tac ctxt n); +*} + method_setup sep_select = {* - Scan.lift Parse.int >> (fn n => fn ctxt => SIMPLE_METHOD' (sep_select_tac ctxt n)) + Scan.lift Parse.int >> sep_select_method *} "Select nth separation conjunct in conclusion" method_setup sep_select_asm = {* - Scan.lift Parse.int >> (fn n => fn ctxt => SIMPLE_METHOD' (sep_select_asm_tac ctxt n)) + Scan.lift Parse.int >> sep_select_asm_method *} "Select nth separation conjunct in assumptions" section {* Substitution *} +ML {* + fun sep_subst_method ctxt occs thms = + SIMPLE_METHOD' (sep_subst_tac ctxt occs thms); + fun sep_subst_asm_method ctxt occs thms = + SIMPLE_METHOD' (sep_subst_asm_tac ctxt occs thms); + + val sep_subst_parser = + Args.mode "asm" + -- Scan.lift (Scan.optional (Args.parens (Scan.repeat Parse.nat)) [0]) + -- Attrib.thms; +*} + method_setup "sep_subst" = {* - Scan.lift (Args.mode "asm" -- Scan.optional (Args.parens (Scan.repeat Parse.nat)) [0]) -- - Attrib.thms >> (fn ((asm, occs), thms) => fn ctxt => - SIMPLE_METHOD' ((if asm then sep_subst_asm_tac else sep_subst_tac) ctxt occs thms)) + sep_subst_parser >> + (fn ((asm, occs), thms) => fn ctxt => + (if asm then sep_subst_asm_method else sep_subst_method) ctxt occs thms) *} "single-step substitution after solving one separation logic assumption" section {* Forward Reasoning *} +ML {* + fun sep_drule_method thms ctxt = SIMPLE_METHOD' (sep_dtac ctxt thms); + fun sep_frule_method thms ctxt = SIMPLE_METHOD' (sep_ftac ctxt thms); +*} + method_setup "sep_drule" = {* - Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' (sep_dtac ctxt thms)) + Attrib.thms >> sep_drule_method *} "drule after separating conjunction reordering" method_setup "sep_frule" = {* - Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' (sep_ftac ctxt thms)) + Attrib.thms >> sep_frule_method *} "frule after separating conjunction reordering" section {* Backward Reasoning *} +ML {* + fun sep_rule_method thms ctxt = SIMPLE_METHOD' (sep_rtac ctxt thms) +*} + method_setup "sep_rule" = {* - Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' (sep_rtac ctxt thms)) + Attrib.thms >> sep_rule_method *} "applies rule after separating conjunction reordering" @@ -97,31 +125,37 @@ fun sep_cancel_smart_tac_rules ctxt etacs = sep_cancel_smart_tac ctxt (FIRST' ([atac] @ etacs)); - val sep_cancel_syntax = Method.sections [ - Args.add -- Args.colon >> K (I, SepCancel_Rules.add)] >> K (); -*} - -method_setup sep_cancel = {* - sep_cancel_syntax >> (fn _ => fn ctxt => + fun sep_cancel_method ctxt = let val etacs = map etac (SepCancel_Rules.get ctxt); in SIMPLE_METHOD' (sep_cancel_smart_tac_rules ctxt etacs) - end) + end; + + val sep_cancel_syntax = Method.sections [ + Args.add -- Args.colon >> K (I, SepCancel_Rules.add)]; +*} + +method_setup sep_cancel = {* + sep_cancel_syntax >> K sep_cancel_method *} "Separating conjunction conjunct cancellation" text {* As above, but use blast with a depth limit to figure out where cancellation can be done. *} -method_setup sep_cancel_blast = {* - sep_cancel_syntax >> (fn _ => fn ctxt => +ML {* + fun sep_cancel_blast_method ctxt = let val rules = SepCancel_Rules.get ctxt; val tac = Blast.depth_tac (ctxt addIs rules) 10; in SIMPLE_METHOD' (sep_cancel_smart_tac ctxt tac) - end) + end; +*} + +method_setup sep_cancel_blast = {* + sep_cancel_syntax >> K sep_cancel_blast_method *} "Separating conjunction conjunct cancellation using blast" end diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/Sep_Tactics.thy~ --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/Sep_Tactics.thy~ Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,161 @@ +(* Authors: Gerwin Klein and Rafal Kolanski, 2012 + Maintainers: Gerwin Klein + Rafal Kolanski +*) + +header "Separation Logic Tactics" + +theory Sep_Tactics +imports Separation_Algebra +begin + +ML_file "sep_tactics.ML" + +text {* A number of proof methods to assist with reasoning about separation logic. *} + + +section {* Selection (move-to-front) tactics *} + +ML {* + fun sep_select_method n ctxt = + Method.SIMPLE_METHOD' (sep_select_tac ctxt n); + fun sep_select_asm_method n ctxt = + Method.SIMPLE_METHOD' (sep_select_asm_tac ctxt n); +*} + +method_setup sep_select = {* + Scan.lift Parse.int >> sep_select_method +*} "Select nth separation conjunct in conclusion" + +method_setup sep_select_asm = {* + Scan.lift Parse.int >> sep_select_asm_method +*} "Select nth separation conjunct in assumptions" + + +section {* Substitution *} + +ML {* + fun sep_subst_method ctxt occs thms = + SIMPLE_METHOD' (sep_subst_tac ctxt occs thms); + fun sep_subst_asm_method ctxt occs thms = + SIMPLE_METHOD' (sep_subst_asm_tac ctxt occs thms); + + val sep_subst_parser = + Args.mode "asm" + -- Scan.lift (Scan.optional (Args.parens (Scan.repeat Parse.nat)) [0]) + -- Attrib.thms; +*} + +method_setup "sep_subst" = {* + sep_subst_parser >> + (fn ((asm, occs), thms) => fn ctxt => + (if asm then sep_subst_asm_method else sep_subst_method) ctxt occs thms) +*} +"single-step substitution after solving one separation logic assumption" + + +section {* Forward Reasoning *} + +ML {* + fun sep_drule_method thms ctxt = SIMPLE_METHOD' (sep_dtac ctxt thms); + fun sep_frule_method thms ctxt = SIMPLE_METHOD' (sep_ftac ctxt thms); +*} + +method_setup "sep_drule" = {* + Attrib.thms >> sep_drule_method +*} "drule after separating conjunction reordering" + +method_setup "sep_frule" = {* + Attrib.thms >> sep_frule_method +*} "frule after separating conjunction reordering" + + +section {* Backward Reasoning *} + +ML {* + fun sep_rule_method thms ctxt = SIMPLE_METHOD' (sep_rtac ctxt thms) +*} + +method_setup "sep_rule" = {* + Attrib.thms >> sep_rule_method +*} "applies rule after separating conjunction reordering" + + +section {* Cancellation of Common Conjuncts via Elimination Rules *} + +ML {* + structure SepCancel_Rules = Named_Thms ( + val name = @{binding "sep_cancel"}; + val description = "sep_cancel rules"; + ); +*} +setup SepCancel_Rules.setup + +text {* + The basic @{text sep_cancel_tac} is minimal. It only eliminates + erule-derivable conjuncts between an assumption and the conclusion. + + To have a more useful tactic, we augment it with more logic, to proceed as + follows: + \begin{itemize} + \item try discharge the goal first using @{text tac} + \item if that fails, invoke @{text sep_cancel_tac} + \item if @{text sep_cancel_tac} succeeds + \begin{itemize} + \item try to finish off with tac (but ok if that fails) + \item try to finish off with @{term sep_true} (but ok if that fails) + \end{itemize} + \end{itemize} + *} + +ML {* + fun sep_cancel_smart_tac ctxt tac = + let fun TRY' tac = tac ORELSE' (K all_tac) + in + tac + ORELSE' (sep_cancel_tac ctxt tac + THEN' TRY' tac + THEN' TRY' (rtac @{thm TrueI})) + ORELSE' (etac @{thm sep_conj_sep_emptyE} + THEN' sep_cancel_tac ctxt tac + THEN' TRY' tac + THEN' TRY' (rtac @{thm TrueI})) + end; + + fun sep_cancel_smart_tac_rules ctxt etacs = + sep_cancel_smart_tac ctxt (FIRST' ([atac] @ etacs)); + + fun sep_cancel_method ctxt = + let + val etacs = map etac (SepCancel_Rules.get ctxt); + in + SIMPLE_METHOD' (sep_cancel_smart_tac_rules ctxt etacs) + end; + + val sep_cancel_syntax = Method.sections [ + Args.add -- Args.colon >> K (I, SepCancel_Rules.add)]; +*} + +method_setup sep_cancel = {* + sep_cancel_syntax >> K sep_cancel_method +*} "Separating conjunction conjunct cancellation" + +text {* + As above, but use blast with a depth limit to figure out where cancellation + can be done. *} + +ML {* + fun sep_cancel_blast_method ctxt = + let + val rules = SepCancel_Rules.get ctxt; + val tac = Blast.depth_tac (ctxt addIs rules) 10; + in + SIMPLE_METHOD' (sep_cancel_smart_tac ctxt tac) + end; +*} + +method_setup sep_cancel_blast = {* + sep_cancel_syntax >> K sep_cancel_blast_method +*} "Separating conjunction conjunct cancellation using blast" + +end diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/Separation_Algebra.thy --- a/Separation_Algebra/Separation_Algebra.thy Fri Sep 12 00:47:15 2014 +0800 +++ b/Separation_Algebra/Separation_Algebra.thy Sat Sep 13 10:07:14 2014 +0800 @@ -35,7 +35,7 @@ pred_K :: "'b \ 'a \ 'b" ("\_\") where "\f\ \ \s. f" -(* was an abbreviation *) +(* abbreviation *) definition pred_ex :: "('b \ 'a \ bool) \ 'a \ bool" (binder "EXS " 10) where "EXS x. P x \ \s. \x. P x s" @@ -216,6 +216,7 @@ then obtain x y z where "P x" and "Q y" and "R z" and "x ## y" and "x ## z" and "y ## z" and "x ## y + z" and "h = x + y + z" + thm sep_disj_addD by (fastforce elim!: sep_conjE simp: sep_add_ac dest: sep_disj_addD) thus "?lhs h" by (metis sep_conj_def sep_disj_addI1) diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/Separation_Algebra.thy-orig --- a/Separation_Algebra/Separation_Algebra.thy-orig Fri Sep 12 00:47:15 2014 +0800 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,815 +0,0 @@ -(* Authors: Gerwin Klein and Rafal Kolanski, 2012 - Maintainers: Gerwin Klein - Rafal Kolanski -*) - -header "Abstract Separation Algebra" - -theory Separation_Algebra -imports Main -begin - - -text {* This theory is the main abstract separation algebra development *} - - -section {* Input syntax for lifting boolean predicates to separation predicates *} - -abbreviation (input) - pred_and :: "('a \ bool) \ ('a \ bool) \ 'a \ bool" (infixr "and" 35) where - "a and b \ \s. a s \ b s" - -abbreviation (input) - pred_or :: "('a \ bool) \ ('a \ bool) \ 'a \ bool" (infixr "or" 30) where - "a or b \ \s. a s \ b s" - -abbreviation (input) - pred_not :: "('a \ bool) \ 'a \ bool" ("not _" [40] 40) where - "not a \ \s. \a s" - -abbreviation (input) - pred_imp :: "('a \ bool) \ ('a \ bool) \ 'a \ bool" (infixr "imp" 25) where - "a imp b \ \s. a s \ b s" - -abbreviation (input) - pred_K :: "'b \ 'a \ 'b" ("\_\") where - "\f\ \ \s. f" - -abbreviation (input) - pred_ex :: "('b \ 'a \ bool) \ 'a \ bool" (binder "EXS " 10) where - "EXS x. P x \ \s. \x. P x s" - -abbreviation (input) - pred_all :: "('b \ 'a \ bool) \ 'a \ bool" (binder "ALLS " 10) where - "ALLS x. P x \ \s. \x. P x s" - - -section {* Associative/Commutative Monoid Basis of Separation Algebras *} - -class pre_sep_algebra = zero + plus + - fixes sep_disj :: "'a => 'a => bool" (infix "##" 60) - - assumes sep_disj_zero [simp]: "x ## 0" - assumes sep_disj_commuteI: "x ## y \ y ## x" - - assumes sep_add_zero [simp]: "x + 0 = x" - assumes sep_add_commute: "x ## y \ x + y = y + x" - - assumes sep_add_assoc: - "\ x ## y; y ## z; x ## z \ \ (x + y) + z = x + (y + z)" -begin - -lemma sep_disj_commute: "x ## y = y ## x" - by (blast intro: sep_disj_commuteI) - -lemma sep_add_left_commute: - assumes a: "a ## b" "b ## c" "a ## c" - shows "b + (a + c) = a + (b + c)" (is "?lhs = ?rhs") -proof - - have "?lhs = b + a + c" using a - by (simp add: sep_add_assoc[symmetric] sep_disj_commute) - also have "... = a + b + c" using a - by (simp add: sep_add_commute sep_disj_commute) - also have "... = ?rhs" using a - by (simp add: sep_add_assoc sep_disj_commute) - finally show ?thesis . -qed - -lemmas sep_add_ac = sep_add_assoc sep_add_commute sep_add_left_commute - sep_disj_commute (* nearly always necessary *) - -end - - -section {* Separation Algebra as Defined by Calcagno et al. *} - -class sep_algebra = pre_sep_algebra + - assumes sep_disj_addD1: "\ x ## y + z; y ## z \ \ x ## y" - assumes sep_disj_addI1: "\ x ## y + z; y ## z \ \ x + y ## z" -begin - -subsection {* Basic Construct Definitions and Abbreviations *} - -definition - sep_conj :: "('a \ bool) \ ('a \ bool) \ ('a \ bool)" (infixr "**" 35) - where - "P ** Q \ \h. \x y. x ## y \ h = x + y \ P x \ Q y" - -notation - sep_conj (infixr "\*" 35) - -definition - sep_empty :: "'a \ bool" ("\") where - "\ \ \h. h = 0" - -definition - sep_impl :: "('a \ bool) \ ('a \ bool) \ ('a \ bool)" (infixr "\*" 25) - where - "P \* Q \ \h. \h'. h ## h' \ P h' \ Q (h + h')" - -definition - sep_substate :: "'a => 'a => bool" (infix "\" 60) where - "x \ y \ \z. x ## z \ x + z = y" - -(* We want these to be abbreviations not definitions, because basic True and - False will occur by simplification in sep_conj terms *) -abbreviation - "sep_true \ \True\" - -abbreviation - "sep_false \ \False\" - -definition - sep_list_conj :: "('a \ bool) list \ ('a \ bool)" ("\* _" [60] 90) where - "sep_list_conj Ps \ foldl (op **) \ Ps" - - -subsection {* Disjunction/Addition Properties *} - -lemma disjoint_zero_sym [simp]: "0 ## x" - by (simp add: sep_disj_commute) - -lemma sep_add_zero_sym [simp]: "0 + x = x" - by (simp add: sep_add_commute) - -lemma sep_disj_addD2: "\ x ## y + z; y ## z \ \ x ## z" - by (metis sep_disj_addD1 sep_add_ac) - -lemma sep_disj_addD: "\ x ## y + z; y ## z \ \ x ## y \ x ## z" - by (metis sep_disj_addD1 sep_disj_addD2) - -lemma sep_add_disjD: "\ x + y ## z; x ## y \ \ x ## z \ y ## z" - by (metis sep_disj_addD sep_disj_commuteI) - -lemma sep_disj_addI2: - "\ x ## y + z; y ## z \ \ x + z ## y" - by (metis sep_add_ac sep_disj_addI1) - -lemma sep_add_disjI1: - "\ x + y ## z; x ## y \ \ x + z ## y" - by (metis sep_add_ac sep_add_disjD sep_disj_addI2) - -lemma sep_add_disjI2: - "\ x + y ## z; x ## y \ \ z + y ## x" - by (metis sep_add_ac sep_add_disjD sep_disj_addI2) - -lemma sep_disj_addI3: - "x + y ## z \ x ## y \ x ## y + z" - by (metis sep_add_ac sep_add_disjD sep_add_disjI2) - -lemma sep_disj_add: - "\ y ## z; x ## y \ \ x ## y + z = x + y ## z" - by (metis sep_disj_addI1 sep_disj_addI3) - - -subsection {* Substate Properties *} - -lemma sep_substate_disj_add: - "x ## y \ x \ x + y" - unfolding sep_substate_def by blast - -lemma sep_substate_disj_add': - "x ## y \ x \ y + x" - by (simp add: sep_add_ac sep_substate_disj_add) - - -subsection {* Separating Conjunction Properties *} - -lemma sep_conjD: - "(P \* Q) h \ \x y. x ## y \ h = x + y \ P x \ Q y" - by (simp add: sep_conj_def) - -lemma sep_conjE: - "\ (P ** Q) h; \x y. \ P x; Q y; x ## y; h = x + y \ \ X \ \ X" - by (auto simp: sep_conj_def) - -lemma sep_conjI: - "\ P x; Q y; x ## y; h = x + y \ \ (P ** Q) h" - by (auto simp: sep_conj_def) - -lemma sep_conj_commuteI: - "(P ** Q) h \ (Q ** P) h" - by (auto intro!: sep_conjI elim!: sep_conjE simp: sep_add_ac) - -lemma sep_conj_commute: - "(P ** Q) = (Q ** P)" - by (rule ext) (auto intro: sep_conj_commuteI) - -lemma sep_conj_assoc: - "((P ** Q) ** R) = (P ** Q ** R)" (is "?lhs = ?rhs") -proof (rule ext, rule iffI) - fix h - assume a: "?lhs h" - then obtain x y z where "P x" and "Q y" and "R z" - and "x ## y" and "x ## z" and "y ## z" and "x + y ## z" - and "h = x + y + z" - by (auto dest!: sep_conjD dest: sep_add_disjD) - moreover - then have "x ## y + z" - by (simp add: sep_disj_add) - ultimately - show "?rhs h" - by (auto simp: sep_add_ac intro!: sep_conjI) -next - fix h - assume a: "?rhs h" - then obtain x y z where "P x" and "Q y" and "R z" - and "x ## y" and "x ## z" and "y ## z" and "x ## y + z" - and "h = x + y + z" - by (fastforce elim!: sep_conjE simp: sep_add_ac dest: sep_disj_addD) - thus "?lhs h" - by (metis sep_conj_def sep_disj_addI1) -qed - -lemma sep_conj_impl: - "\ (P ** Q) h; \h. P h \ P' h; \h. Q h \ Q' h \ \ (P' ** Q') h" - by (erule sep_conjE, auto intro!: sep_conjI) - -lemma sep_conj_impl1: - assumes P: "\h. P h \ I h" - shows "(P ** R) h \ (I ** R) h" - by (auto intro: sep_conj_impl P) - -lemma sep_globalise: - "\ (P ** R) h; (\h. P h \ Q h) \ \ (Q ** R) h" - by (fast elim: sep_conj_impl) - -lemma sep_conj_trivial_strip2: - "Q = R \ (Q ** P) = (R ** P)" by simp - -lemma disjoint_subheaps_exist: - "\x y. x ## y \ h = x + y" - by (rule_tac x=0 in exI, auto) - -lemma sep_conj_left_commute: (* for permutative rewriting *) - "(P ** (Q ** R)) = (Q ** (P ** R))" (is "?x = ?y") -proof - - have "?x = ((Q ** R) ** P)" by (simp add: sep_conj_commute) - also have "\ = (Q ** (R ** P))" by (subst sep_conj_assoc, simp) - finally show ?thesis by (simp add: sep_conj_commute) -qed - -lemmas sep_conj_ac = sep_conj_commute sep_conj_assoc sep_conj_left_commute - -lemma ab_semigroup_mult_sep_conj: "class.ab_semigroup_mult op **" - by (unfold_locales) - (auto simp: sep_conj_ac) - -lemma sep_empty_zero [simp,intro!]: "\ 0" - by (simp add: sep_empty_def) - - -subsection {* Properties of @{text sep_true} and @{text sep_false} *} - -lemma sep_conj_sep_true: - "P h \ (P ** sep_true) h" - by (simp add: sep_conjI[where y=0]) - -lemma sep_conj_sep_true': - "P h \ (sep_true ** P) h" - by (simp add: sep_conjI[where x=0]) - -lemma sep_conj_true [simp]: - "(sep_true ** sep_true) = sep_true" - unfolding sep_conj_def - by (auto intro!: ext intro: disjoint_subheaps_exist) - -lemma sep_conj_false_right [simp]: - "(P ** sep_false) = sep_false" - by (force elim: sep_conjE intro!: ext) - -lemma sep_conj_false_left [simp]: - "(sep_false ** P) = sep_false" - by (subst sep_conj_commute) (rule sep_conj_false_right) - - - -subsection {* Properties of zero (@{const sep_empty}) *} - -lemma sep_conj_empty [simp]: - "(P ** \) = P" - by (simp add: sep_conj_def sep_empty_def) - -lemma sep_conj_empty'[simp]: - "(\ ** P) = P" - by (subst sep_conj_commute, rule sep_conj_empty) - -lemma sep_conj_sep_emptyI: - "P h \ (P ** \) h" - by simp - -lemma sep_conj_sep_emptyE: - "\ P s; (P ** \) s \ (Q ** R) s \ \ (Q ** R) s" - by simp - -lemma monoid_add: "class.monoid_add (op **) \" - by (unfold_locales) (auto simp: sep_conj_ac) - -lemma comm_monoid_add: "class.comm_monoid_add op ** \" - by (unfold_locales) (auto simp: sep_conj_ac) - - -subsection {* Properties of top (@{text sep_true}) *} - -lemma sep_conj_true_P [simp]: - "(sep_true ** (sep_true ** P)) = (sep_true ** P)" - by (simp add: sep_conj_assoc[symmetric]) - -lemma sep_conj_disj: - "((P or Q) ** R) = ((P ** R) or (Q ** R))" - by (auto simp: sep_conj_def intro!: ext) - -lemma sep_conj_sep_true_left: - "(P ** Q) h \ (sep_true ** Q) h" - by (erule sep_conj_impl, simp+) - -lemma sep_conj_sep_true_right: - "(P ** Q) h \ (P ** sep_true) h" - by (subst (asm) sep_conj_commute, drule sep_conj_sep_true_left, - simp add: sep_conj_ac) - - -subsection {* Separating Conjunction with Quantifiers *} - -lemma sep_conj_conj: - "((P and Q) ** R) h \ ((P ** R) and (Q ** R)) h" - by (force intro: sep_conjI elim!: sep_conjE) - -lemma sep_conj_exists1: - "((EXS x. P x) ** Q) = (EXS x. (P x ** Q))" - by (force intro!: ext intro: sep_conjI elim: sep_conjE) - -lemma sep_conj_exists2: - "(P ** (EXS x. Q x)) = (EXS x. P ** Q x)" - by (force intro!: sep_conjI ext elim!: sep_conjE) - -lemmas sep_conj_exists = sep_conj_exists1 sep_conj_exists2 - -lemma sep_conj_spec: - "((ALLS x. P x) ** Q) h \ (P x ** Q) h" - by (force intro: sep_conjI elim: sep_conjE) - - -subsection {* Properties of Separating Implication *} - -lemma sep_implI: - assumes a: "\h'. \ h ## h'; P h' \ \ Q (h + h')" - shows "(P \* Q) h" - unfolding sep_impl_def by (auto elim: a) - -lemma sep_implD: - "(x \* y) h \ \h'. h ## h' \ x h' \ y (h + h')" - by (force simp: sep_impl_def) - -lemma sep_implE: - "(x \* y) h \ (\h'. h ## h' \ x h' \ y (h + h') \ Q) \ Q" - by (auto dest: sep_implD) - -lemma sep_impl_sep_true [simp]: - "(P \* sep_true) = sep_true" - by (force intro!: sep_implI ext) - -lemma sep_impl_sep_false [simp]: - "(sep_false \* P) = sep_true" - by (force intro!: sep_implI ext) - -lemma sep_impl_sep_true_P: - "(sep_true \* P) h \ P h" - by (clarsimp dest!: sep_implD elim!: allE[where x=0]) - -lemma sep_impl_sep_true_false [simp]: - "(sep_true \* sep_false) = sep_false" - by (force intro!: ext dest: sep_impl_sep_true_P) - -lemma sep_conj_sep_impl: - "\ P h; \h. (P ** Q) h \ R h \ \ (Q \* R) h" -proof (rule sep_implI) - fix h' h - assume "P h" and "h ## h'" and "Q h'" - hence "(P ** Q) (h + h')" by (force intro: sep_conjI) - moreover assume "\h. (P ** Q) h \ R h" - ultimately show "R (h + h')" by simp -qed - -lemma sep_conj_sep_impl2: - "\ (P ** Q) h; \h. P h \ (Q \* R) h \ \ R h" - by (force dest: sep_implD elim: sep_conjE) - -lemma sep_conj_sep_impl_sep_conj2: - "(P ** R) h \ (P ** (Q \* (Q ** R))) h" - by (erule (1) sep_conj_impl, erule sep_conj_sep_impl, simp add: sep_conj_ac) - - -subsection {* Pure assertions *} - -definition - pure :: "('a \ bool) \ bool" where - "pure P \ \h h'. P h = P h'" - -lemma pure_sep_true: - "pure sep_true" - by (simp add: pure_def) - -lemma pure_sep_false: - "pure sep_true" - by (simp add: pure_def) - -lemma pure_split: - "pure P = (P = sep_true \ P = sep_false)" - by (force simp: pure_def intro!: ext) - -lemma pure_sep_conj: - "\ pure P; pure Q \ \ pure (P \* Q)" - by (force simp: pure_split) - -lemma pure_sep_impl: - "\ pure P; pure Q \ \ pure (P \* Q)" - by (force simp: pure_split) - -lemma pure_conj_sep_conj: - "\ (P and Q) h; pure P \ pure Q \ \ (P \* Q) h" - by (metis pure_def sep_add_zero sep_conjI sep_conj_commute sep_disj_zero) - -lemma pure_sep_conj_conj: - "\ (P \* Q) h; pure P; pure Q \ \ (P and Q) h" - by (force simp: pure_split) - -lemma pure_conj_sep_conj_assoc: - "pure P \ ((P and Q) \* R) = (P and (Q \* R))" - by (auto simp: pure_split) - -lemma pure_sep_impl_impl: - "\ (P \* Q) h; pure P \ \ P h \ Q h" - by (force simp: pure_split dest: sep_impl_sep_true_P) - -lemma pure_impl_sep_impl: - "\ P h \ Q h; pure P; pure Q \ \ (P \* Q) h" - by (force simp: pure_split) - -lemma pure_conj_right: "(Q \* (\P'\ and Q')) = (\P'\ and (Q \* Q'))" - by (rule ext, rule, rule, clarsimp elim!: sep_conjE) - (erule sep_conj_impl, auto) - -lemma pure_conj_right': "(Q \* (P' and \Q'\)) = (\Q'\ and (Q \* P'))" - by (simp add: conj_comms pure_conj_right) - -lemma pure_conj_left: "((\P'\ and Q') \* Q) = (\P'\ and (Q' \* Q))" - by (simp add: pure_conj_right sep_conj_ac) - -lemma pure_conj_left': "((P' and \Q'\) \* Q) = (\Q'\ and (P' \* Q))" - by (subst conj_comms, subst pure_conj_left, simp) - -lemmas pure_conj = pure_conj_right pure_conj_right' pure_conj_left - pure_conj_left' - -declare pure_conj[simp add] - - -subsection {* Intuitionistic assertions *} - -definition intuitionistic :: "('a \ bool) \ bool" where - "intuitionistic P \ \h h'. P h \ h \ h' \ P h'" - -lemma intuitionisticI: - "(\h h'. \ P h; h \ h' \ \ P h') \ intuitionistic P" - by (unfold intuitionistic_def, fast) - -lemma intuitionisticD: - "\ intuitionistic P; P h; h \ h' \ \ P h'" - by (unfold intuitionistic_def, fast) - -lemma pure_intuitionistic: - "pure P \ intuitionistic P" - by (clarsimp simp: intuitionistic_def pure_def, fast) - -lemma intuitionistic_conj: - "\ intuitionistic P; intuitionistic Q \ \ intuitionistic (P and Q)" - by (force intro: intuitionisticI dest: intuitionisticD) - -lemma intuitionistic_disj: - "\ intuitionistic P; intuitionistic Q \ \ intuitionistic (P or Q)" - by (force intro: intuitionisticI dest: intuitionisticD) - -lemma intuitionistic_forall: - "(\x. intuitionistic (P x)) \ intuitionistic (ALLS x. P x)" - by (force intro: intuitionisticI dest: intuitionisticD) - -lemma intuitionistic_exists: - "(\x. intuitionistic (P x)) \ intuitionistic (EXS x. P x)" - by (force intro: intuitionisticI dest: intuitionisticD) - -lemma intuitionistic_sep_conj_sep_true: - "intuitionistic (sep_true \* P)" -proof (rule intuitionisticI) - fix h h' r - assume a: "(sep_true \* P) h" - then obtain x y where P: "P y" and h: "h = x + y" and xyd: "x ## y" - by - (drule sep_conjD, clarsimp) - moreover assume a2: "h \ h'" - then obtain z where h': "h' = h + z" and hzd: "h ## z" - by (clarsimp simp: sep_substate_def) - - moreover have "(P \* sep_true) (y + (x + z))" - using P h hzd xyd - by (metis sep_add_disjI1 sep_disj_commute sep_conjI) - ultimately show "(sep_true \* P) h'" using hzd - by (auto simp: sep_conj_commute sep_add_ac dest!: sep_disj_addD) -qed - -lemma intuitionistic_sep_impl_sep_true: - "intuitionistic (sep_true \* P)" -proof (rule intuitionisticI) - fix h h' - assume imp: "(sep_true \* P) h" and hh': "h \ h'" - - from hh' obtain z where h': "h' = h + z" and hzd: "h ## z" - by (clarsimp simp: sep_substate_def) - show "(sep_true \* P) h'" using imp h' hzd - apply (clarsimp dest!: sep_implD) - apply (metis sep_add_assoc sep_add_disjD sep_disj_addI3 sep_implI) - done -qed - -lemma intuitionistic_sep_conj: - assumes ip: "intuitionistic (P::('a \ bool))" - shows "intuitionistic (P \* Q)" -proof (rule intuitionisticI) - fix h h' - assume sc: "(P \* Q) h" and hh': "h \ h'" - - from hh' obtain z where h': "h' = h + z" and hzd: "h ## z" - by (clarsimp simp: sep_substate_def) - - from sc obtain x y where px: "P x" and qy: "Q y" - and h: "h = x + y" and xyd: "x ## y" - by (clarsimp simp: sep_conj_def) - - have "x ## z" using hzd h xyd - by (metis sep_add_disjD) - - with ip px have "P (x + z)" - by (fastforce elim: intuitionisticD sep_substate_disj_add) - - thus "(P \* Q) h'" using h' h hzd qy xyd - by (metis (full_types) sep_add_commute sep_add_disjD sep_add_disjI2 - sep_add_left_commute sep_conjI) -qed - -lemma intuitionistic_sep_impl: - assumes iq: "intuitionistic Q" - shows "intuitionistic (P \* Q)" -proof (rule intuitionisticI) - fix h h' - assume imp: "(P \* Q) h" and hh': "h \ h'" - - from hh' obtain z where h': "h' = h + z" and hzd: "h ## z" - by (clarsimp simp: sep_substate_def) - - { - fix x - assume px: "P x" and hzx: "h + z ## x" - - have "h + x \ h + x + z" using hzx hzd - by (metis sep_add_disjI1 sep_substate_def) - - with imp hzd iq px hzx - have "Q (h + z + x)" - by (metis intuitionisticD sep_add_assoc sep_add_ac sep_add_disjD sep_implE) - } - - with imp h' hzd iq show "(P \* Q) h'" - by (fastforce intro: sep_implI) -qed - -lemma strongest_intuitionistic: - "\ (\Q. (\h. (Q h \ (P \* sep_true) h)) \ intuitionistic Q \ - Q \ (P \* sep_true) \ (\h. P h \ Q h))" - by (fastforce intro!: ext sep_substate_disj_add - dest!: sep_conjD intuitionisticD) - -lemma weakest_intuitionistic: - "\ (\Q. (\h. ((sep_true \* P) h \ Q h)) \ intuitionistic Q \ - Q \ (sep_true \* P) \ (\h. Q h \ P h))" - apply (clarsimp intro!: ext) - apply (rule iffI) - apply (rule sep_implI) - apply (drule_tac h="x" and h'="x + h'" in intuitionisticD) - apply (clarsimp simp: sep_add_ac sep_substate_disj_add)+ - done - -lemma intuitionistic_sep_conj_sep_true_P: - "\ (P \* sep_true) s; intuitionistic P \ \ P s" - by (force dest: intuitionisticD elim: sep_conjE sep_substate_disj_add) - -lemma intuitionistic_sep_conj_sep_true_simp: - "intuitionistic P \ (P \* sep_true) = P" - by (fast intro!: sep_conj_sep_true ext - elim: intuitionistic_sep_conj_sep_true_P) - -lemma intuitionistic_sep_impl_sep_true_P: - "\ P h; intuitionistic P \ \ (sep_true \* P) h" - by (force intro!: sep_implI dest: intuitionisticD - intro: sep_substate_disj_add) - -lemma intuitionistic_sep_impl_sep_true_simp: - "intuitionistic P \ (sep_true \* P) = P" - by (fast intro!: ext - elim: sep_impl_sep_true_P intuitionistic_sep_impl_sep_true_P) - - -subsection {* Strictly exact assertions *} - -definition strictly_exact :: "('a \ bool) \ bool" where - "strictly_exact P \ \h h'. P h \ P h' \ h = h'" - -lemma strictly_exactD: - "\ strictly_exact P; P h; P h' \ \ h = h'" - by (unfold strictly_exact_def, fast) - -lemma strictly_exactI: - "(\h h'. \ P h; P h' \ \ h = h') \ strictly_exact P" - by (unfold strictly_exact_def, fast) - -lemma strictly_exact_sep_conj: - "\ strictly_exact P; strictly_exact Q \ \ strictly_exact (P \* Q)" - apply (rule strictly_exactI) - apply (erule sep_conjE)+ - apply (drule_tac h="x" and h'="xa" in strictly_exactD, assumption+) - apply (drule_tac h="y" and h'="ya" in strictly_exactD, assumption+) - apply clarsimp - done - -lemma strictly_exact_conj_impl: - "\ (Q \* sep_true) h; P h; strictly_exact Q \ \ (Q \* (Q \* P)) h" - by (force intro: sep_conjI sep_implI dest: strictly_exactD elim!: sep_conjE - simp: sep_add_commute sep_add_assoc) - -end - -interpretation sep: ab_semigroup_mult "op **" - by (rule ab_semigroup_mult_sep_conj) - -interpretation sep: comm_monoid_add "op **" \ - by (rule comm_monoid_add) - - -section {* Separation Algebra with Stronger, but More Intuitive Disjunction Axiom *} - -class stronger_sep_algebra = pre_sep_algebra + - assumes sep_add_disj_eq [simp]: "y ## z \ x ## y + z = (x ## y \ x ## z)" -begin - -lemma sep_disj_add_eq [simp]: "x ## y \ x + y ## z = (x ## z \ y ## z)" - by (metis sep_add_disj_eq sep_disj_commute) - -subclass sep_algebra by default auto - -end - - -section {* Folding separating conjunction over lists of predicates *} - -lemma sep_list_conj_Nil [simp]: "\* [] = \" - by (simp add: sep_list_conj_def) - -(* apparently these two are rarely used and had to be removed from List.thy *) -lemma (in semigroup_add) foldl_assoc: -shows "foldl op+ (x+y) zs = x + (foldl op+ y zs)" -by (induct zs arbitrary: y) (simp_all add:add_assoc) - -lemma (in monoid_add) foldl_absorb0: -shows "x + (foldl op+ 0 zs) = foldl op+ x zs" -by (induct zs) (simp_all add:foldl_assoc) - -lemma sep_list_conj_Cons [simp]: "\* (x#xs) = (x ** \* xs)" - by (simp add: sep_list_conj_def sep.foldl_absorb0) - -lemma sep_list_conj_append [simp]: "\* (xs @ ys) = (\* xs ** \* ys)" - by (simp add: sep_list_conj_def sep.foldl_absorb0) - -lemma (in comm_monoid_add) foldl_map_filter: - "foldl op + 0 (map f (filter P xs)) + - foldl op + 0 (map f (filter (not P) xs)) - = foldl op + 0 (map f xs)" -proof (induct xs) - case Nil thus ?case by clarsimp -next - case (Cons x xs) - hence IH: "foldl op + 0 (map f xs) = - foldl op + 0 (map f (filter P xs)) + - foldl op + 0 (map f [x\xs . \ P x])" - by (simp only: eq_commute) - - have foldl_Cons': - "\x xs. foldl op + 0 (x # xs) = x + (foldl op + 0 xs)" - by (simp, subst foldl_absorb0[symmetric], rule refl) - - { assume "P x" - hence ?case by (auto simp del: foldl_Cons simp add: foldl_Cons' IH add_ac) - } moreover { - assume "\ P x" - hence ?case by (auto simp del: foldl_Cons simp add: foldl_Cons' IH add_ac) - } - ultimately show ?case by blast -qed - - -section {* Separation Algebra with a Cancellative Monoid (for completeness) *} - -text {* - Separation algebra with a cancellative monoid. The results of being a precise - assertion (distributivity over separating conjunction) require this. - although we never actually use this property in our developments, we keep - it here for completeness. - *} -class cancellative_sep_algebra = sep_algebra + - assumes sep_add_cancelD: "\ x + z = y + z ; x ## z ; y ## z \ \ x = y" -begin - -definition - (* In any heap, there exists at most one subheap for which P holds *) - precise :: "('a \ bool) \ bool" where - "precise P = (\h hp hp'. hp \ h \ P hp \ hp' \ h \ P hp' \ hp = hp')" - -lemma "precise (op = s)" - by (metis (full_types) precise_def) - -lemma sep_add_cancel: - "x ## z \ y ## z \ (x + z = y + z) = (x = y)" - by (metis sep_add_cancelD) - -lemma precise_distribute: - "precise P = (\Q R. ((Q and R) \* P) = ((Q \* P) and (R \* P)))" -proof (rule iffI) - assume pp: "precise P" - { - fix Q R - fix h hp hp' s - - { assume a: "((Q and R) \* P) s" - hence "((Q \* P) and (R \* P)) s" - by (fastforce dest!: sep_conjD elim: sep_conjI) - } - moreover - { assume qs: "(Q \* P) s" and qr: "(R \* P) s" - - from qs obtain x y where sxy: "s = x + y" and xy: "x ## y" - and x: "Q x" and y: "P y" - by (fastforce dest!: sep_conjD) - from qr obtain x' y' where sxy': "s = x' + y'" and xy': "x' ## y'" - and x': "R x'" and y': "P y'" - by (fastforce dest!: sep_conjD) - - from sxy have ys: "y \ x + y" using xy - by (fastforce simp: sep_substate_disj_add' sep_disj_commute) - from sxy' have ys': "y' \ x' + y'" using xy' - by (fastforce simp: sep_substate_disj_add' sep_disj_commute) - - from pp have yy: "y = y'" using sxy sxy' xy xy' y y' ys ys' - by (fastforce simp: precise_def) - - hence "x = x'" using sxy sxy' xy xy' - by (fastforce dest!: sep_add_cancelD) - - hence "((Q and R) \* P) s" using sxy x x' yy y' xy' - by (fastforce intro: sep_conjI) - } - ultimately - have "((Q and R) \* P) s = ((Q \* P) and (R \* P)) s" using pp by blast - } - thus "\Q R. ((Q and R) \* P) = ((Q \* P) and (R \* P))" by (blast intro!: ext) - -next - assume a: "\Q R. ((Q and R) \* P) = ((Q \* P) and (R \* P))" - thus "precise P" - proof (clarsimp simp: precise_def) - fix h hp hp' Q R - assume hp: "hp \ h" and hp': "hp' \ h" and php: "P hp" and php': "P hp'" - - obtain z where hhp: "h = hp + z" and hpz: "hp ## z" using hp - by (clarsimp simp: sep_substate_def) - obtain z' where hhp': "h = hp' + z'" and hpz': "hp' ## z'" using hp' - by (clarsimp simp: sep_substate_def) - - have h_eq: "z' + hp' = z + hp" using hhp hhp' hpz hpz' - by (fastforce simp: sep_add_ac) - - from hhp hhp' a hpz hpz' h_eq - have "\Q R. ((Q and R) \* P) (z + hp) = ((Q \* P) and (R \* P)) (z' + hp')" - by (fastforce simp: h_eq sep_add_ac sep_conj_commute) - - hence "((op = z and op = z') \* P) (z + hp) = - ((op = z \* P) and (op = z' \* P)) (z' + hp')" by blast - - thus "hp = hp'" using php php' hpz hpz' h_eq - by (fastforce dest!: iffD2 cong: conj_cong - simp: sep_add_ac sep_add_cancel sep_conj_def) - qed -qed - -lemma strictly_precise: "strictly_exact P \ precise P" - by (metis precise_def strictly_exactD) - -end - -end diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/Separation_Algebra.thy~ --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/Separation_Algebra.thy~ Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,815 @@ +(* Authors: Gerwin Klein and Rafal Kolanski, 2012 + Maintainers: Gerwin Klein + Rafal Kolanski +*) + +header "Abstract Separation Algebra" + +theory Separation_Algebra +imports Main +begin + + +text {* This theory is the main abstract separation algebra development *} + + +section {* Input syntax for lifting boolean predicates to separation predicates *} + +abbreviation (input) + pred_and :: "('a \ bool) \ ('a \ bool) \ 'a \ bool" (infixr "and" 35) where + "a and b \ \s. a s \ b s" + +abbreviation (input) + pred_or :: "('a \ bool) \ ('a \ bool) \ 'a \ bool" (infixr "or" 30) where + "a or b \ \s. a s \ b s" + +abbreviation (input) + pred_not :: "('a \ bool) \ 'a \ bool" ("not _" [40] 40) where + "not a \ \s. \a s" + +abbreviation (input) + pred_imp :: "('a \ bool) \ ('a \ bool) \ 'a \ bool" (infixr "imp" 25) where + "a imp b \ \s. a s \ b s" + +abbreviation (input) + pred_K :: "'b \ 'a \ 'b" ("\_\") where + "\f\ \ \s. f" + +(* abbreviation *) + definition pred_ex :: "('b \ 'a \ bool) \ 'a \ bool" (binder "EXS " 10) where + "EXS x. P x \ \s. \x. P x s" + +abbreviation (input) + pred_all :: "('b \ 'a \ bool) \ 'a \ bool" (binder "ALLS " 10) where + "ALLS x. P x \ \s. \x. P x s" + + +section {* Associative/Commutative Monoid Basis of Separation Algebras *} + +class pre_sep_algebra = zero + plus + + fixes sep_disj :: "'a => 'a => bool" (infix "##" 60) + + assumes sep_disj_zero [simp]: "x ## 0" + assumes sep_disj_commuteI: "x ## y \ y ## x" + + assumes sep_add_zero [simp]: "x + 0 = x" + assumes sep_add_commute: "x ## y \ x + y = y + x" + + assumes sep_add_assoc: + "\ x ## y; y ## z; x ## z \ \ (x + y) + z = x + (y + z)" +begin + +lemma sep_disj_commute: "x ## y = y ## x" + by (blast intro: sep_disj_commuteI) + +lemma sep_add_left_commute: + assumes a: "a ## b" "b ## c" "a ## c" + shows "b + (a + c) = a + (b + c)" (is "?lhs = ?rhs") +proof - + have "?lhs = b + a + c" using a + by (simp add: sep_add_assoc[symmetric] sep_disj_commute) + also have "... = a + b + c" using a + by (simp add: sep_add_commute sep_disj_commute) + also have "... = ?rhs" using a + by (simp add: sep_add_assoc sep_disj_commute) + finally show ?thesis . +qed + +lemmas sep_add_ac = sep_add_assoc sep_add_commute sep_add_left_commute + sep_disj_commute (* nearly always necessary *) + +end + + +section {* Separation Algebra as Defined by Calcagno et al. *} + +class sep_algebra = pre_sep_algebra + + assumes sep_disj_addD1: "\ x ## y + z; y ## z \ \ x ## y" + assumes sep_disj_addI1: "\ x ## y + z; y ## z \ \ x + y ## z" +begin + +subsection {* Basic Construct Definitions and Abbreviations *} + +definition + sep_conj :: "('a \ bool) \ ('a \ bool) \ ('a \ bool)" (infixr "**" 35) + where + "P ** Q \ \h. \x y. x ## y \ h = x + y \ P x \ Q y" + +notation + sep_conj (infixr "\*" 35) + +definition + sep_empty :: "'a \ bool" ("\") where + "\ \ \h. h = 0" + +definition + sep_impl :: "('a \ bool) \ ('a \ bool) \ ('a \ bool)" (infixr "\*" 25) + where + "P \* Q \ \h. \h'. h ## h' \ P h' \ Q (h + h')" + +definition + sep_substate :: "'a => 'a => bool" (infix "\" 60) where + "x \ y \ \z. x ## z \ x + z = y" + +(* We want these to be abbreviations not definitions, because basic True and + False will occur by simplification in sep_conj terms *) +abbreviation + "sep_true \ \True\" + +abbreviation + "sep_false \ \False\" + +definition + sep_list_conj :: "('a \ bool) list \ ('a \ bool)" ("\* _" [60] 90) where + "sep_list_conj Ps \ foldl (op **) \ Ps" + + +subsection {* Disjunction/Addition Properties *} + +lemma disjoint_zero_sym [simp]: "0 ## x" + by (simp add: sep_disj_commute) + +lemma sep_add_zero_sym [simp]: "0 + x = x" + by (simp add: sep_add_commute) + +lemma sep_disj_addD2: "\ x ## y + z; y ## z \ \ x ## z" + by (metis sep_disj_addD1 sep_add_ac) + +lemma sep_disj_addD: "\ x ## y + z; y ## z \ \ x ## y \ x ## z" + by (metis sep_disj_addD1 sep_disj_addD2) + +lemma sep_add_disjD: "\ x + y ## z; x ## y \ \ x ## z \ y ## z" + by (metis sep_disj_addD sep_disj_commuteI) + +lemma sep_disj_addI2: + "\ x ## y + z; y ## z \ \ x + z ## y" + by (metis sep_add_ac sep_disj_addI1) + +lemma sep_add_disjI1: + "\ x + y ## z; x ## y \ \ x + z ## y" + by (metis sep_add_ac sep_add_disjD sep_disj_addI2) + +lemma sep_add_disjI2: + "\ x + y ## z; x ## y \ \ z + y ## x" + by (metis sep_add_ac sep_add_disjD sep_disj_addI2) + +lemma sep_disj_addI3: + "x + y ## z \ x ## y \ x ## y + z" + by (metis sep_add_ac sep_add_disjD sep_add_disjI2) + +lemma sep_disj_add: + "\ y ## z; x ## y \ \ x ## y + z = x + y ## z" + by (metis sep_disj_addI1 sep_disj_addI3) + + +subsection {* Substate Properties *} + +lemma sep_substate_disj_add: + "x ## y \ x \ x + y" + unfolding sep_substate_def by blast + +lemma sep_substate_disj_add': + "x ## y \ x \ y + x" + by (simp add: sep_add_ac sep_substate_disj_add) + + +subsection {* Separating Conjunction Properties *} + +lemma sep_conjD: + "(P \* Q) h \ \x y. x ## y \ h = x + y \ P x \ Q y" + by (simp add: sep_conj_def) + +lemma sep_conjE: + "\ (P ** Q) h; \x y. \ P x; Q y; x ## y; h = x + y \ \ X \ \ X" + by (auto simp: sep_conj_def) + +lemma sep_conjI: + "\ P x; Q y; x ## y; h = x + y \ \ (P ** Q) h" + by (auto simp: sep_conj_def) + +lemma sep_conj_commuteI: + "(P ** Q) h \ (Q ** P) h" + by (auto intro!: sep_conjI elim!: sep_conjE simp: sep_add_ac) + +lemma sep_conj_commute: + "(P ** Q) = (Q ** P)" + by (rule ext) (auto intro: sep_conj_commuteI) + +lemma sep_conj_assoc: + "((P ** Q) ** R) = (P ** Q ** R)" (is "?lhs = ?rhs") +proof (rule ext, rule iffI) + fix h + assume a: "?lhs h" + then obtain x y z where "P x" and "Q y" and "R z" + and "x ## y" and "x ## z" and "y ## z" and "x + y ## z" + and "h = x + y + z" + by (auto dest!: sep_conjD dest: sep_add_disjD) + moreover + then have "x ## y + z" + by (simp add: sep_disj_add) + ultimately + show "?rhs h" + by (auto simp: sep_add_ac intro!: sep_conjI) +next + fix h + assume a: "?rhs h" + then obtain x y z where "P x" and "Q y" and "R z" + and "x ## y" and "x ## z" and "y ## z" and "x ## y + z" + and "h = x + y + z" + by (fastforce elim!: sep_conjE simp: sep_add_ac dest: sep_disj_addD) + thus "?lhs h" + by (metis sep_conj_def sep_disj_addI1) +qed + +lemma sep_conj_impl: + "\ (P ** Q) h; \h. P h \ P' h; \h. Q h \ Q' h \ \ (P' ** Q') h" + by (erule sep_conjE, auto intro!: sep_conjI) + +lemma sep_conj_impl1: + assumes P: "\h. P h \ I h" + shows "(P ** R) h \ (I ** R) h" + by (auto intro: sep_conj_impl P) + +lemma sep_globalise: + "\ (P ** R) h; (\h. P h \ Q h) \ \ (Q ** R) h" + by (fast elim: sep_conj_impl) + +lemma sep_conj_trivial_strip2: + "Q = R \ (Q ** P) = (R ** P)" by simp + +lemma disjoint_subheaps_exist: + "\x y. x ## y \ h = x + y" + by (rule_tac x=0 in exI, auto) + +lemma sep_conj_left_commute: (* for permutative rewriting *) + "(P ** (Q ** R)) = (Q ** (P ** R))" (is "?x = ?y") +proof - + have "?x = ((Q ** R) ** P)" by (simp add: sep_conj_commute) + also have "\ = (Q ** (R ** P))" by (subst sep_conj_assoc, simp) + finally show ?thesis by (simp add: sep_conj_commute) +qed + +lemmas sep_conj_ac = sep_conj_commute sep_conj_assoc sep_conj_left_commute + +lemma ab_semigroup_mult_sep_conj: "class.ab_semigroup_mult op **" + by (unfold_locales) + (auto simp: sep_conj_ac) + +lemma sep_empty_zero [simp,intro!]: "\ 0" + by (simp add: sep_empty_def) + + +subsection {* Properties of @{text sep_true} and @{text sep_false} *} + +lemma sep_conj_sep_true: + "P h \ (P ** sep_true) h" + by (simp add: sep_conjI[where y=0]) + +lemma sep_conj_sep_true': + "P h \ (sep_true ** P) h" + by (simp add: sep_conjI[where x=0]) + +lemma sep_conj_true [simp]: + "(sep_true ** sep_true) = sep_true" + unfolding sep_conj_def + by (auto intro!: ext intro: disjoint_subheaps_exist) + +lemma sep_conj_false_right [simp]: + "(P ** sep_false) = sep_false" + by (force elim: sep_conjE intro!: ext) + +lemma sep_conj_false_left [simp]: + "(sep_false ** P) = sep_false" + by (subst sep_conj_commute) (rule sep_conj_false_right) + + + +subsection {* Properties of zero (@{const sep_empty}) *} + +lemma sep_conj_empty [simp]: + "(P ** \) = P" + by (simp add: sep_conj_def sep_empty_def) + +lemma sep_conj_empty'[simp]: + "(\ ** P) = P" + by (subst sep_conj_commute, rule sep_conj_empty) + +lemma sep_conj_sep_emptyI: + "P h \ (P ** \) h" + by simp + +lemma sep_conj_sep_emptyE: + "\ P s; (P ** \) s \ (Q ** R) s \ \ (Q ** R) s" + by simp + +lemma monoid_add: "class.monoid_add (op **) \" + by (unfold_locales) (auto simp: sep_conj_ac) + +lemma comm_monoid_add: "class.comm_monoid_add op ** \" + by (unfold_locales) (auto simp: sep_conj_ac) + + +subsection {* Properties of top (@{text sep_true}) *} + +lemma sep_conj_true_P [simp]: + "(sep_true ** (sep_true ** P)) = (sep_true ** P)" + by (simp add: sep_conj_assoc[symmetric]) + +lemma sep_conj_disj: + "((P or Q) ** R) = ((P ** R) or (Q ** R))" + by (auto simp: sep_conj_def intro!: ext) + +lemma sep_conj_sep_true_left: + "(P ** Q) h \ (sep_true ** Q) h" + by (erule sep_conj_impl, simp+) + +lemma sep_conj_sep_true_right: + "(P ** Q) h \ (P ** sep_true) h" + by (subst (asm) sep_conj_commute, drule sep_conj_sep_true_left, + simp add: sep_conj_ac) + + +subsection {* Separating Conjunction with Quantifiers *} + +lemma sep_conj_conj: + "((P and Q) ** R) h \ ((P ** R) and (Q ** R)) h" + by (force intro: sep_conjI elim!: sep_conjE) + +lemma sep_conj_exists1: + "((EXS x. P x) ** Q) = (EXS x. (P x ** Q))" + by (unfold pred_ex_def, force intro!: ext intro: sep_conjI elim: sep_conjE) + +lemma sep_conj_exists2: + "(P ** (EXS x. Q x)) = (EXS x. P ** Q x)" + by (unfold pred_ex_def, force intro!: sep_conjI ext elim!: sep_conjE) + +lemmas sep_conj_exists = sep_conj_exists1 sep_conj_exists2 + +lemma sep_conj_spec: + "((ALLS x. P x) ** Q) h \ (P x ** Q) h" + by (force intro: sep_conjI elim: sep_conjE) + + +subsection {* Properties of Separating Implication *} + +lemma sep_implI: + assumes a: "\h'. \ h ## h'; P h' \ \ Q (h + h')" + shows "(P \* Q) h" + unfolding sep_impl_def by (auto elim: a) + +lemma sep_implD: + "(x \* y) h \ \h'. h ## h' \ x h' \ y (h + h')" + by (force simp: sep_impl_def) + +lemma sep_implE: + "(x \* y) h \ (\h'. h ## h' \ x h' \ y (h + h') \ Q) \ Q" + by (auto dest: sep_implD) + +lemma sep_impl_sep_true [simp]: + "(P \* sep_true) = sep_true" + by (force intro!: sep_implI ext) + +lemma sep_impl_sep_false [simp]: + "(sep_false \* P) = sep_true" + by (force intro!: sep_implI ext) + +lemma sep_impl_sep_true_P: + "(sep_true \* P) h \ P h" + by (clarsimp dest!: sep_implD elim!: allE[where x=0]) + +lemma sep_impl_sep_true_false [simp]: + "(sep_true \* sep_false) = sep_false" + by (force intro!: ext dest: sep_impl_sep_true_P) + +lemma sep_conj_sep_impl: + "\ P h; \h. (P ** Q) h \ R h \ \ (Q \* R) h" +proof (rule sep_implI) + fix h' h + assume "P h" and "h ## h'" and "Q h'" + hence "(P ** Q) (h + h')" by (force intro: sep_conjI) + moreover assume "\h. (P ** Q) h \ R h" + ultimately show "R (h + h')" by simp +qed + +lemma sep_conj_sep_impl2: + "\ (P ** Q) h; \h. P h \ (Q \* R) h \ \ R h" + by (force dest: sep_implD elim: sep_conjE) + +lemma sep_conj_sep_impl_sep_conj2: + "(P ** R) h \ (P ** (Q \* (Q ** R))) h" + by (erule (1) sep_conj_impl, erule sep_conj_sep_impl, simp add: sep_conj_ac) + + +subsection {* Pure assertions *} + +definition + pure :: "('a \ bool) \ bool" where + "pure P \ \h h'. P h = P h'" + +lemma pure_sep_true: + "pure sep_true" + by (simp add: pure_def) + +lemma pure_sep_false: + "pure sep_true" + by (simp add: pure_def) + +lemma pure_split: + "pure P = (P = sep_true \ P = sep_false)" + by (force simp: pure_def intro!: ext) + +lemma pure_sep_conj: + "\ pure P; pure Q \ \ pure (P \* Q)" + by (force simp: pure_split) + +lemma pure_sep_impl: + "\ pure P; pure Q \ \ pure (P \* Q)" + by (force simp: pure_split) + +lemma pure_conj_sep_conj: + "\ (P and Q) h; pure P \ pure Q \ \ (P \* Q) h" + by (metis pure_def sep_add_zero sep_conjI sep_conj_commute sep_disj_zero) + +lemma pure_sep_conj_conj: + "\ (P \* Q) h; pure P; pure Q \ \ (P and Q) h" + by (force simp: pure_split) + +lemma pure_conj_sep_conj_assoc: + "pure P \ ((P and Q) \* R) = (P and (Q \* R))" + by (auto simp: pure_split) + +lemma pure_sep_impl_impl: + "\ (P \* Q) h; pure P \ \ P h \ Q h" + by (force simp: pure_split dest: sep_impl_sep_true_P) + +lemma pure_impl_sep_impl: + "\ P h \ Q h; pure P; pure Q \ \ (P \* Q) h" + by (force simp: pure_split) + +lemma pure_conj_right: "(Q \* (\P'\ and Q')) = (\P'\ and (Q \* Q'))" + by (rule ext, rule, rule, clarsimp elim!: sep_conjE) + (erule sep_conj_impl, auto) + +lemma pure_conj_right': "(Q \* (P' and \Q'\)) = (\Q'\ and (Q \* P'))" + by (simp add: conj_comms pure_conj_right) + +lemma pure_conj_left: "((\P'\ and Q') \* Q) = (\P'\ and (Q' \* Q))" + by (simp add: pure_conj_right sep_conj_ac) + +lemma pure_conj_left': "((P' and \Q'\) \* Q) = (\Q'\ and (P' \* Q))" + by (subst conj_comms, subst pure_conj_left, simp) + +lemmas pure_conj = pure_conj_right pure_conj_right' pure_conj_left + pure_conj_left' + +declare pure_conj[simp add] + + +subsection {* Intuitionistic assertions *} + +definition intuitionistic :: "('a \ bool) \ bool" where + "intuitionistic P \ \h h'. P h \ h \ h' \ P h'" + +lemma intuitionisticI: + "(\h h'. \ P h; h \ h' \ \ P h') \ intuitionistic P" + by (unfold intuitionistic_def, fast) + +lemma intuitionisticD: + "\ intuitionistic P; P h; h \ h' \ \ P h'" + by (unfold intuitionistic_def, fast) + +lemma pure_intuitionistic: + "pure P \ intuitionistic P" + by (clarsimp simp: intuitionistic_def pure_def, fast) + +lemma intuitionistic_conj: + "\ intuitionistic P; intuitionistic Q \ \ intuitionistic (P and Q)" + by (force intro: intuitionisticI dest: intuitionisticD) + +lemma intuitionistic_disj: + "\ intuitionistic P; intuitionistic Q \ \ intuitionistic (P or Q)" + by (force intro: intuitionisticI dest: intuitionisticD) + +lemma intuitionistic_forall: + "(\x. intuitionistic (P x)) \ intuitionistic (ALLS x. P x)" + by (force intro: intuitionisticI dest: intuitionisticD) + +lemma intuitionistic_exists: + "(\x. intuitionistic (P x)) \ intuitionistic (EXS x. P x)" + by (unfold pred_ex_def, force intro: intuitionisticI dest: intuitionisticD) + +lemma intuitionistic_sep_conj_sep_true: + "intuitionistic (sep_true \* P)" +proof (rule intuitionisticI) + fix h h' r + assume a: "(sep_true \* P) h" + then obtain x y where P: "P y" and h: "h = x + y" and xyd: "x ## y" + by - (drule sep_conjD, clarsimp) + moreover assume a2: "h \ h'" + then obtain z where h': "h' = h + z" and hzd: "h ## z" + by (clarsimp simp: sep_substate_def) + + moreover have "(P \* sep_true) (y + (x + z))" + using P h hzd xyd + by (metis sep_add_disjI1 sep_disj_commute sep_conjI) + ultimately show "(sep_true \* P) h'" using hzd + by (auto simp: sep_conj_commute sep_add_ac dest!: sep_disj_addD) +qed + +lemma intuitionistic_sep_impl_sep_true: + "intuitionistic (sep_true \* P)" +proof (rule intuitionisticI) + fix h h' + assume imp: "(sep_true \* P) h" and hh': "h \ h'" + + from hh' obtain z where h': "h' = h + z" and hzd: "h ## z" + by (clarsimp simp: sep_substate_def) + show "(sep_true \* P) h'" using imp h' hzd + apply (clarsimp dest!: sep_implD) + apply (metis sep_add_assoc sep_add_disjD sep_disj_addI3 sep_implI) + done +qed + +lemma intuitionistic_sep_conj: + assumes ip: "intuitionistic (P::('a \ bool))" + shows "intuitionistic (P \* Q)" +proof (rule intuitionisticI) + fix h h' + assume sc: "(P \* Q) h" and hh': "h \ h'" + + from hh' obtain z where h': "h' = h + z" and hzd: "h ## z" + by (clarsimp simp: sep_substate_def) + + from sc obtain x y where px: "P x" and qy: "Q y" + and h: "h = x + y" and xyd: "x ## y" + by (clarsimp simp: sep_conj_def) + + have "x ## z" using hzd h xyd + by (metis sep_add_disjD) + + with ip px have "P (x + z)" + by (fastforce elim: intuitionisticD sep_substate_disj_add) + + thus "(P \* Q) h'" using h' h hzd qy xyd + by (metis (full_types) sep_add_commute sep_add_disjD sep_add_disjI2 + sep_add_left_commute sep_conjI) +qed + +lemma intuitionistic_sep_impl: + assumes iq: "intuitionistic Q" + shows "intuitionistic (P \* Q)" +proof (rule intuitionisticI) + fix h h' + assume imp: "(P \* Q) h" and hh': "h \ h'" + + from hh' obtain z where h': "h' = h + z" and hzd: "h ## z" + by (clarsimp simp: sep_substate_def) + + { + fix x + assume px: "P x" and hzx: "h + z ## x" + + have "h + x \ h + x + z" using hzx hzd + by (metis sep_add_disjI1 sep_substate_def) + + with imp hzd iq px hzx + have "Q (h + z + x)" + by (metis intuitionisticD sep_add_assoc sep_add_ac sep_add_disjD sep_implE) + } + + with imp h' hzd iq show "(P \* Q) h'" + by (fastforce intro: sep_implI) +qed + +lemma strongest_intuitionistic: + "\ (\Q. (\h. (Q h \ (P \* sep_true) h)) \ intuitionistic Q \ + Q \ (P \* sep_true) \ (\h. P h \ Q h))" + by (fastforce intro!: ext sep_substate_disj_add + dest!: sep_conjD intuitionisticD) + +lemma weakest_intuitionistic: + "\ (\Q. (\h. ((sep_true \* P) h \ Q h)) \ intuitionistic Q \ + Q \ (sep_true \* P) \ (\h. Q h \ P h))" + apply (clarsimp intro!: ext) + apply (rule iffI) + apply (rule sep_implI) + apply (drule_tac h="x" and h'="x + h'" in intuitionisticD) + apply (clarsimp simp: sep_add_ac sep_substate_disj_add)+ + done + +lemma intuitionistic_sep_conj_sep_true_P: + "\ (P \* sep_true) s; intuitionistic P \ \ P s" + by (force dest: intuitionisticD elim: sep_conjE sep_substate_disj_add) + +lemma intuitionistic_sep_conj_sep_true_simp: + "intuitionistic P \ (P \* sep_true) = P" + by (fast intro!: sep_conj_sep_true ext + elim: intuitionistic_sep_conj_sep_true_P) + +lemma intuitionistic_sep_impl_sep_true_P: + "\ P h; intuitionistic P \ \ (sep_true \* P) h" + by (force intro!: sep_implI dest: intuitionisticD + intro: sep_substate_disj_add) + +lemma intuitionistic_sep_impl_sep_true_simp: + "intuitionistic P \ (sep_true \* P) = P" + by (fast intro!: ext + elim: sep_impl_sep_true_P intuitionistic_sep_impl_sep_true_P) + + +subsection {* Strictly exact assertions *} + +definition strictly_exact :: "('a \ bool) \ bool" where + "strictly_exact P \ \h h'. P h \ P h' \ h = h'" + +lemma strictly_exactD: + "\ strictly_exact P; P h; P h' \ \ h = h'" + by (unfold strictly_exact_def, fast) + +lemma strictly_exactI: + "(\h h'. \ P h; P h' \ \ h = h') \ strictly_exact P" + by (unfold strictly_exact_def, fast) + +lemma strictly_exact_sep_conj: + "\ strictly_exact P; strictly_exact Q \ \ strictly_exact (P \* Q)" + apply (rule strictly_exactI) + apply (erule sep_conjE)+ + apply (drule_tac h="x" and h'="xa" in strictly_exactD, assumption+) + apply (drule_tac h="y" and h'="ya" in strictly_exactD, assumption+) + apply clarsimp + done + +lemma strictly_exact_conj_impl: + "\ (Q \* sep_true) h; P h; strictly_exact Q \ \ (Q \* (Q \* P)) h" + by (force intro: sep_conjI sep_implI dest: strictly_exactD elim!: sep_conjE + simp: sep_add_commute sep_add_assoc) + +end + +interpretation sep: ab_semigroup_mult "op **" + by (rule ab_semigroup_mult_sep_conj) + +interpretation sep: comm_monoid_add "op **" \ + by (rule comm_monoid_add) + + +section {* Separation Algebra with Stronger, but More Intuitive Disjunction Axiom *} + +class stronger_sep_algebra = pre_sep_algebra + + assumes sep_add_disj_eq [simp]: "y ## z \ x ## y + z = (x ## y \ x ## z)" +begin + +lemma sep_disj_add_eq [simp]: "x ## y \ x + y ## z = (x ## z \ y ## z)" + by (metis sep_add_disj_eq sep_disj_commute) + +subclass sep_algebra by default auto + +end + + +section {* Folding separating conjunction over lists of predicates *} + +lemma sep_list_conj_Nil [simp]: "\* [] = \" + by (simp add: sep_list_conj_def) + +(* apparently these two are rarely used and had to be removed from List.thy *) +lemma (in semigroup_add) foldl_assoc: +shows "foldl op+ (x+y) zs = x + (foldl op+ y zs)" +by (induct zs arbitrary: y) (simp_all add:add_assoc) + +lemma (in monoid_add) foldl_absorb0: +shows "x + (foldl op+ 0 zs) = foldl op+ x zs" +by (induct zs) (simp_all add:foldl_assoc) + +lemma sep_list_conj_Cons [simp]: "\* (x#xs) = (x ** \* xs)" + by (simp add: sep_list_conj_def sep.foldl_absorb0) + +lemma sep_list_conj_append [simp]: "\* (xs @ ys) = (\* xs ** \* ys)" + by (simp add: sep_list_conj_def sep.foldl_absorb0) + +lemma (in comm_monoid_add) foldl_map_filter: + "foldl op + 0 (map f (filter P xs)) + + foldl op + 0 (map f (filter (not P) xs)) + = foldl op + 0 (map f xs)" +proof (induct xs) + case Nil thus ?case by clarsimp +next + case (Cons x xs) + hence IH: "foldl op + 0 (map f xs) = + foldl op + 0 (map f (filter P xs)) + + foldl op + 0 (map f [x\xs . \ P x])" + by (simp only: eq_commute) + + have foldl_Cons': + "\x xs. foldl op + 0 (x # xs) = x + (foldl op + 0 xs)" + by (simp, subst foldl_absorb0[symmetric], rule refl) + + { assume "P x" + hence ?case by (auto simp del: foldl_Cons simp add: foldl_Cons' IH add_ac) + } moreover { + assume "\ P x" + hence ?case by (auto simp del: foldl_Cons simp add: foldl_Cons' IH add_ac) + } + ultimately show ?case by blast +qed + + +section {* Separation Algebra with a Cancellative Monoid (for completeness) *} + +text {* + Separation algebra with a cancellative monoid. The results of being a precise + assertion (distributivity over separating conjunction) require this. + although we never actually use this property in our developments, we keep + it here for completeness. + *} +class cancellative_sep_algebra = sep_algebra + + assumes sep_add_cancelD: "\ x + z = y + z ; x ## z ; y ## z \ \ x = y" +begin + +definition + (* In any heap, there exists at most one subheap for which P holds *) + precise :: "('a \ bool) \ bool" where + "precise P = (\h hp hp'. hp \ h \ P hp \ hp' \ h \ P hp' \ hp = hp')" + +lemma "precise (op = s)" + by (metis (full_types) precise_def) + +lemma sep_add_cancel: + "x ## z \ y ## z \ (x + z = y + z) = (x = y)" + by (metis sep_add_cancelD) + +lemma precise_distribute: + "precise P = (\Q R. ((Q and R) \* P) = ((Q \* P) and (R \* P)))" +proof (rule iffI) + assume pp: "precise P" + { + fix Q R + fix h hp hp' s + + { assume a: "((Q and R) \* P) s" + hence "((Q \* P) and (R \* P)) s" + by (fastforce dest!: sep_conjD elim: sep_conjI) + } + moreover + { assume qs: "(Q \* P) s" and qr: "(R \* P) s" + + from qs obtain x y where sxy: "s = x + y" and xy: "x ## y" + and x: "Q x" and y: "P y" + by (fastforce dest!: sep_conjD) + from qr obtain x' y' where sxy': "s = x' + y'" and xy': "x' ## y'" + and x': "R x'" and y': "P y'" + by (fastforce dest!: sep_conjD) + + from sxy have ys: "y \ x + y" using xy + by (fastforce simp: sep_substate_disj_add' sep_disj_commute) + from sxy' have ys': "y' \ x' + y'" using xy' + by (fastforce simp: sep_substate_disj_add' sep_disj_commute) + + from pp have yy: "y = y'" using sxy sxy' xy xy' y y' ys ys' + by (fastforce simp: precise_def) + + hence "x = x'" using sxy sxy' xy xy' + by (fastforce dest!: sep_add_cancelD) + + hence "((Q and R) \* P) s" using sxy x x' yy y' xy' + by (fastforce intro: sep_conjI) + } + ultimately + have "((Q and R) \* P) s = ((Q \* P) and (R \* P)) s" using pp by blast + } + thus "\Q R. ((Q and R) \* P) = ((Q \* P) and (R \* P))" by (blast intro!: ext) + +next + assume a: "\Q R. ((Q and R) \* P) = ((Q \* P) and (R \* P))" + thus "precise P" + proof (clarsimp simp: precise_def) + fix h hp hp' Q R + assume hp: "hp \ h" and hp': "hp' \ h" and php: "P hp" and php': "P hp'" + + obtain z where hhp: "h = hp + z" and hpz: "hp ## z" using hp + by (clarsimp simp: sep_substate_def) + obtain z' where hhp': "h = hp' + z'" and hpz': "hp' ## z'" using hp' + by (clarsimp simp: sep_substate_def) + + have h_eq: "z' + hp' = z + hp" using hhp hhp' hpz hpz' + by (fastforce simp: sep_add_ac) + + from hhp hhp' a hpz hpz' h_eq + have "\Q R. ((Q and R) \* P) (z + hp) = ((Q \* P) and (R \* P)) (z' + hp')" + by (fastforce simp: h_eq sep_add_ac sep_conj_commute) + + hence "((op = z and op = z') \* P) (z + hp) = + ((op = z \* P) and (op = z' \* P)) (z' + hp')" by blast + + thus "hp = hp'" using php php' hpz hpz' h_eq + by (fastforce dest!: iffD2 cong: conj_cong + simp: sep_add_ac sep_add_cancel sep_conj_def) + qed +qed + +lemma strictly_precise: "strictly_exact P \ precise P" + by (metis precise_def strictly_exactD) + +end + +end diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/Separation_Algebra_Alt.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/Separation_Algebra_Alt.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,313 @@ +(* Author: Gerwin Klein, 2012 + Maintainers: Gerwin Klein + Rafal Kolanski +*) + +header "Abstract Separation Logic, Alternative Definition" + +theory Separation_Algebra_Alt +imports Main +begin + +text {* + This theory contains an alternative definition of speration algebra, + following Calcagno et al very closely. While some of the abstract + development is more algebraic, it is cumbersome to instantiate. + We only use it to prove equivalence and to give an impression of how + it would look like. +*} + +(* The @{text "++"} notation is a horrible choice, but this theory is + only intended to show how the development would look like, not to + actually use it. We remove the notation for map-add so it doesn't + conflict. +*) +no_notation map_add (infixl "++" 100) + +definition + lift2 :: "('a => 'b => 'c option) \ 'a option => 'b option => 'c option" +where + "lift2 f a b \ case (a,b) of (Some a, Some b) \ f a b | _ \ None" + + +class sep_algebra_alt = zero + + fixes add :: "'a => 'a => 'a option" (infixr "\" 65) + + assumes add_zero [simp]: "x \ 0 = Some x" + assumes add_comm: "x \ y = y \ x" + assumes add_assoc: "lift2 add a (lift2 add b c) = lift2 add (lift2 add a b) c" + +begin + +definition + disjoint :: "'a => 'a => bool" (infix "##" 60) +where + "a ## b \ a \ b \ None" + +lemma disj_com: "x ## y = y ## x" + by (auto simp: disjoint_def add_comm) + +lemma disj_zero [simp]: "x ## 0" + by (auto simp: disjoint_def) + +lemma disj_zero2 [simp]: "0 ## x" + by (subst disj_com) simp + +lemma add_zero2 [simp]: "0 \ x = Some x" + by (subst add_comm) auto + +definition + substate :: "'a => 'a => bool" (infix "\" 60) where + "a \ b \ \c. a \ c = Some b" + +definition + sep_conj :: "('a \ bool) \ ('a \ bool) \ ('a \ bool)" (infixl "**" 61) + where + "P ** Q \ \s. \p q. p \ q = Some s \ P p \ Q q" + +definition emp :: "'a \ bool" ("\") where + "\ \ \s. s = 0" + +definition + sep_impl :: "('a \ bool) \ ('a \ bool) \ ('a \ bool)" (infixr "\\<^sup>*" 25) + where + "P \\<^sup>* Q \ \h. \h' h''. h \ h' = Some h'' \ P h' \ Q h''" + +definition + "sep_true \ \s. True" + +definition + "sep_false \ \s. False" + + +abbreviation + add2 :: "'a option => 'a option => 'a option" (infixr "++" 65) +where + "add2 == lift2 add" + + +lemma add2_comm: + "a ++ b = b ++ a" + by (simp add: lift2_def add_comm split: option.splits) + +lemma add2_None [simp]: + "x ++ None = None" + by (simp add: lift2_def split: option.splits) + +lemma None_add2 [simp]: + "None ++ x = None" + by (simp add: lift2_def split: option.splits) + +lemma add2_Some_Some: + "Some x ++ Some y = x \ y" + by (simp add: lift2_def) + +lemma add2_zero [simp]: + "Some x ++ Some 0 = Some x" + by (simp add: add2_Some_Some) + +lemma zero_add2 [simp]: + "Some 0 ++ Some x = Some x" + by (simp add: add2_Some_Some) + + +lemma sep_conjE: + "\ (P ** Q) s; \p q. \ P p; Q q; p \ q = Some s \ \ X \ \ X" + by (auto simp: sep_conj_def) + +lemma sep_conjI: + "\ P p; Q q; p \ q = Some s \ \ (P ** Q) s" + by (auto simp: sep_conj_def) + +lemma sep_conj_comI: + "(P ** Q) s \ (Q ** P) s" + by (auto intro!: sep_conjI elim!: sep_conjE simp: add_comm) + +lemma sep_conj_com: + "P ** Q = Q ** P" + by (auto intro: sep_conj_comI intro!: ext) + +lemma lift_to_add2: + "\z \ q = Some s; x \ y = Some q\ \ Some z ++ Some x ++ Some y = Some s" + by (simp add: add2_Some_Some) + +lemma lift_to_add2': + "\q \ z = Some s; x \ y = Some q\ \ (Some x ++ Some y) ++ Some z = Some s" + by (simp add: add2_Some_Some) + +lemma add2_Some: + "(x ++ Some y = Some z) = (\x'. x = Some x' \ x' \ y = Some z)" + by (simp add: lift2_def split: option.splits) + +lemma Some_add2: + "(Some x ++ y = Some z) = (\y'. y = Some y' \ x \ y' = Some z)" + by (simp add: lift2_def split: option.splits) + +lemma sep_conj_assoc: + "P ** (Q ** R) = (P ** Q) ** R" + unfolding sep_conj_def + apply (rule ext) + apply (rule iffI) + apply clarsimp + apply (drule (1) lift_to_add2) + apply (subst (asm) add_assoc) + apply (fastforce simp: add2_Some_Some add2_Some) + apply clarsimp + apply (drule (1) lift_to_add2') + apply (subst (asm) add_assoc [symmetric]) + apply (fastforce simp: add2_Some_Some Some_add2) + done + +lemma sep_true[simp]: "sep_true s" by (simp add: sep_true_def) +lemma sep_false[simp]: "\sep_false x" by (simp add: sep_false_def) + +lemma sep_conj_sep_true: + "P s \ (P ** sep_true) s" + by (auto simp: sep_conjI [where q=0]) + +lemma sep_conj_sep_true': + "P s \ (sep_true ** P) s" + by (auto simp: sep_conjI [where p=0]) + +lemma disjoint_submaps_exist: + "\h\<^isub>0 h\<^isub>1. h\<^isub>0 \ h\<^isub>1 = Some h" + by (rule_tac x=0 in exI, auto) + +lemma sep_conj_true[simp]: + "(sep_true ** sep_true) = sep_true" + unfolding sep_conj_def + by (auto intro!: ext intro: disjoint_submaps_exist) + +lemma sep_conj_false_right[simp]: + "(P ** sep_false) = sep_false" + by (force elim: sep_conjE intro!: ext) + +lemma sep_conj_false_left[simp]: + "(sep_false ** P) = sep_false" + by (subst sep_conj_com) (rule sep_conj_false_right) + +lemma sep_conj_left_com: + "(P ** (Q ** R)) = (Q ** (P ** R))" (is "?x = ?y") +proof - + have "?x = ((Q ** R) ** P)" by (simp add: sep_conj_com) + also have "\ = (Q ** (R ** P))" by (subst sep_conj_assoc, simp) + finally show ?thesis by (simp add: sep_conj_com) +qed + +lemmas sep_conj_ac = sep_conj_com sep_conj_assoc sep_conj_left_com + +lemma empty_empty[simp]: "\ 0" + by (simp add: emp_def) + +lemma sep_conj_empty[simp]: + "(P ** \) = P" + by (simp add: sep_conj_def emp_def) + + lemma sep_conj_empty'[simp]: + "(\ ** P) = P" + by (subst sep_conj_com, rule sep_conj_empty) + +lemma sep_conj_sep_emptyI: + "P s \ (P ** \) s" + by simp + +lemma sep_conj_true_P[simp]: + "(sep_true ** (sep_true ** P)) = (sep_true ** P)" + by (simp add: sep_conj_assoc) + +lemma sep_conj_disj: + "((\s. P s \ Q s) ** R) s = ((P ** R) s \ (Q ** R) s)" (is "?x = (?y \ ?z)") + by (auto simp: sep_conj_def) + +lemma sep_conj_conj: + "((\s. P s \ Q s) ** R) s \ (P ** R) s \ (Q ** R) s" + by (force intro: sep_conjI elim!: sep_conjE) + +lemma sep_conj_exists1: + "((\s. \x. P x s) ** Q) s = (\x. (P x ** Q) s)" + by (force intro: sep_conjI elim: sep_conjE) + +lemma sep_conj_exists2: + "(P ** (\s. \x. Q x s)) = (\s. (\x. (P ** Q x) s))" + by (force intro!: sep_conjI ext elim!: sep_conjE) + +lemmas sep_conj_exists = sep_conj_exists1 sep_conj_exists2 + +lemma sep_conj_forall: + "((\s. \x. P x s) ** Q) s \ (P x ** Q) s" + by (force intro: sep_conjI elim: sep_conjE) + +lemma sep_conj_impl: + "\ (P ** Q) s; \s. P s \ P' s; \s. Q s \ Q' s \ \ (P' ** Q') s" + by (erule sep_conjE, auto intro!: sep_conjI) + +lemma sep_conj_impl1: + assumes P: "\s. P s \ I s" + shows "(P ** R) s \ (I ** R) s" + by (auto intro: sep_conj_impl P) + +lemma sep_conj_sep_true_left: + "(P ** Q) s \ (sep_true ** Q) s" + by (erule sep_conj_impl, simp+) + +lemma sep_conj_sep_true_right: + "(P ** Q) s \ (P ** sep_true) s" + by (subst (asm) sep_conj_com, drule sep_conj_sep_true_left, + simp add: sep_conj_ac) + +lemma sep_globalise: + "\ (P ** R) s; (\s. P s \ Q s) \ \ (Q ** R) s" + by (fast elim: sep_conj_impl) + +lemma sep_implI: + assumes a: "\h' h''. \ h \ h' = Some h''; P h' \ \ Q h''" + shows "(P \\<^sup>* Q) h" + unfolding sep_impl_def by (auto elim: a) + +lemma sep_implD: + "(x \\<^sup>* y) h \ \h' h''. h \ h' = Some h'' \ x h' \ y h''" + by (force simp: sep_impl_def) + +lemma sep_impl_sep_true[simp]: + "(P \\<^sup>* sep_true) = sep_true" + by (force intro!: sep_implI ext) + +lemma sep_impl_sep_false[simp]: + "(sep_false \\<^sup>* P) = sep_true" + by (force intro!: sep_implI ext) + +lemma sep_impl_sep_true_P: + "(sep_true \\<^sup>* P) s \ P s" + apply (drule sep_implD) + apply (erule_tac x=0 in allE) + apply simp + done + +lemma sep_impl_sep_true_false[simp]: + "(sep_true \\<^sup>* sep_false) = sep_false" + by (force intro!: ext dest: sep_impl_sep_true_P) + +lemma sep_conj_sep_impl: + "\ P s; \s. (P ** Q) s \ R s \ \ (Q \\<^sup>* R) s" +proof (rule sep_implI) + fix h' h h'' + assume "P h" and "h \ h' = Some h''" and "Q h'" + hence "(P ** Q) h''" by (force intro: sep_conjI) + moreover assume "\s. (P ** Q) s \ R s" + ultimately show "R h''" by simp +qed + +lemma sep_conj_sep_impl2: + "\ (P ** Q) s; \s. P s \ (Q \\<^sup>* R) s \ \ R s" + by (force dest: sep_implD elim: sep_conjE) + +lemma sep_conj_sep_impl_sep_conj2: + "(P ** R) s \ (P ** (Q \\<^sup>* (Q ** R))) s" + by (erule (1) sep_conj_impl, erule sep_conj_sep_impl, simp add: sep_conj_ac) + +lemma sep_conj_triv_strip2: + "Q = R \ (Q ** P) = (R ** P)" by simp + +end + +end diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/config --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/config Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,9 @@ +# -*- shell-script -*- + +# Get email when automated build fails. May be empty. +# values: "email1 email2 .. emailn" +NOTIFY="kleing@cse.unsw.edu.au rafal.kolanski@nicta.com.au" + +# Participate in frequent (nightly) build (only for small submissions) +# values: "yes" "no" +FREQUENT="yes" diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/document/root.bib --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/document/root.bib Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,6 @@ +@inproceedings{KleinKB-ITP12, +author={Gerwin Klein and Rafal Kolanski and Andrew Boyton}, +title={Mechanised Separation Algebra (Rough Diamond)}, +booktitle={Interactive Theorem Proving (ITP 2012)}, +editor={Beringer and Felty}, +publisher={Springer},series={LNCS},volume={},pages={},year=2012} diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/document/root.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/document/root.tex Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,67 @@ +% $Id: root.tex,v 1.2 2007-11-12 21:00:45 nipkow Exp $ + +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} + +% further packages required for unusual symbols (see also isabellesym.sty) +% use only when needed +\usepackage{amssymb} % for \, \, \, + % \, \, \, + % \, \, \, + % \, \, + % \, \, \ +%\usepackage[greek,english]{babel} % greek for \, + % english for \, + % \ + % default language = last +%\usepackage[latin1]{inputenc} % for \, \, + % \, \, + % \, \, + % \ +%\usepackage[only,bigsqcap]{stmaryrd} % for \ +%\usepackage{eufrak} % for \ ... \, \ ... \ + % (only needed if amssymb not used) +%\usepackage{textcomp} % for \, \ + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + + +\begin{document} + +\title{Separation Algebra} +\author{Gerwin Klein and Rafal Kolanski and Andrew Boyton} +\maketitle + +\begin{abstract} +We present a generic type class implementation of separation +algebra for Isabelle/HOL as well as lemmas and generic tactics which can +be used directly for any instantiation of the type class. + +The ex directory contains example instantiations that include structures +such as a heap or virtual memory. + +The abstract separation algebra is based upon ``Abstract Separation +Logic'' by Calcagno et al. These theories are also the basis of +``Mechanised Separation Algebra'' by the authors \cite{KleinKB-ITP12}. + +The aim of this work is to support and significantly reduce the effort +for future separation logic developments in Isabelle/HOL by factoring +out the part of separation logic that can be treated abstractly once +and for all. This includes developing typical default rule sets for +reasoning as well as automated tactic support for separation logic. +\end{abstract} + +\tableofcontents + +% include generated text of all theories +\input{session} + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/ex/Sep_Tactics_Test.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/ex/Sep_Tactics_Test.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,123 @@ +(* Authors: Gerwin Klein and Rafal Kolanski, 2012 + Maintainers: Gerwin Klein + Rafal Kolanski +*) + +theory Sep_Tactics_Test +imports "../Sep_Tactics" +begin + +text {* Substitution and forward/backward reasoning *} + +typedecl p +typedecl val +typedecl heap + +arities heap :: sep_algebra + +axiomatization + points_to :: "p \ val \ heap \ bool" and + val :: "heap \ p \ val" +where + points_to: "(points_to p v ** P) h \ val h p = v" + + +lemma + "\ Q2 (val h p); (K ** T ** blub ** P ** points_to p v ** P ** J) h \ + \ Q (val h p) (val h p)" + apply (sep_subst (2) points_to) + apply (sep_subst (asm) points_to) + apply (sep_subst points_to) + oops + +lemma + "\ Q2 (val h p); (K ** T ** blub ** P ** points_to p v ** P ** J) h \ + \ Q (val h p) (val h p)" + apply (sep_drule points_to) + apply simp + oops + +lemma + "\ Q2 (val h p); (K ** T ** blub ** P ** points_to p v ** P ** J) h \ + \ Q (val h p) (val h p)" + apply (sep_frule points_to) + apply simp + oops + +consts + update :: "p \ val \ heap \ heap" + +schematic_lemma + assumes a: "\P. (stuff p ** P) H \ (other_stuff p v ** P) (update p v H)" + shows "(X ** Y ** other_stuff p ?v) (update p v H)" + apply (sep_rule a) + oops + + +text {* Example of low-level rewrites *} + +lemma "\ unrelated s ; (P ** Q ** R) s \ \ (A ** B ** Q ** P) s" + apply (tactic {* dtac (mk_sep_select_rule @{context} true (3,1)) 1 *}) + apply (tactic {* rtac (mk_sep_select_rule @{context} false (4,2)) 1 *}) + (* now sep_conj_impl1 can be used *) + apply (erule (1) sep_conj_impl) + oops + + +text {* Conjunct selection *} + +lemma "(A ** B ** Q ** P) s" + apply (sep_select 1) + apply (sep_select 3) + apply (sep_select 4) + oops + +lemma "\ also unrelated; (A ** B ** Q ** P) s \ \ unrelated" + apply (sep_select_asm 2) + oops + + +section {* Test cases for @{text sep_cancel}. *} + +lemma + assumes forward: "\s g p v. A g p v s \ AA g p s " + shows "\xv yv P s y x s. (A g x yv ** A g y yv ** P) s \ (AA g y ** sep_true) s" + by (sep_cancel add: forward) + +lemma + assumes forward: "\s. generic s \ instance s" + shows "(A ** generic ** B) s \ ((instance ** B)**A) s" + by (sep_cancel add: forward)+ + +lemma "\ (A ** B) sa ; (A ** Y) s \ \ (A ** X) s" + apply (sep_cancel) + oops + +lemma "\ (A ** B) sa ; (A ** Y) s \ \ (\s. (A ** X) s) s" + apply (sep_cancel) + oops + +schematic_lemma tt: "\ (B ** A ** C) s \ \ (\s. (A ** ?X) s) s" + by (sep_cancel) + +(* test backtracking on premises with same state *) + +lemma + assumes forward: "\s. generic s \ instance s" + shows "\ (A ** B) s ; (generic ** Y) s \ \ (X ** instance) s" + apply (sep_cancel add: forward) + oops + +lemma + assumes forward: "\s. generic s \ instance s" + shows "generic s \ instance s" + by (sep_cancel add: forward) + +lemma + assumes forward: "\s. generic s \ instance s" + assumes forward2: "\s. instance s \ instance2 s" + shows "generic s \ (instance2 ** sep_true) s" + by (sep_cancel_blast add: forward forward2) + +end + diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/ex/Sep_Tactics_Test.thy~ --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/ex/Sep_Tactics_Test.thy~ Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,122 @@ +(* Authors: Gerwin Klein and Rafal Kolanski, 2012 + Maintainers: Gerwin Klein + Rafal Kolanski +*) + +theory Sep_Tactics_Test +imports "../Sep_Tactics" +begin + +text {* Substitution and forward/backward reasoning *} + +typedecl p +typedecl val +typedecl heap + +arities heap :: sep_algebra + +axiomatization + points_to :: "p \ val \ heap \ bool" and + val :: "heap \ p \ val" +where + points_to: "(points_to p v ** P) h \ val h p = v" + + +lemma + "\ Q2 (val h p); (K ** T ** blub ** P ** points_to p v ** P ** J) h \ + \ Q (val h p) (val h p)" + apply (sep_subst (2) points_to) + apply (sep_subst (asm) points_to) + apply (sep_subst points_to) + oops + +lemma + "\ Q2 (val h p); (K ** T ** blub ** P ** points_to p v ** P ** J) h \ + \ Q (val h p) (val h p)" + apply (sep_drule points_to) + apply simp + oops + +lemma + "\ Q2 (val h p); (K ** T ** blub ** P ** points_to p v ** P ** J) h \ + \ Q (val h p) (val h p)" + apply (sep_frule points_to) + apply simp + oops + +consts + update :: "p \ val \ heap \ heap" + +schematic_lemma + assumes a: "\P. (stuff p ** P) H \ (other_stuff p v ** P) (update p v H)" + shows "(X ** Y ** other_stuff p ?v) (update p v H)" + apply (sep_rule a) + oops + + +text {* Example of low-level rewrites *} + +lemma "\ unrelated s ; (P ** Q ** R) s \ \ (A ** B ** Q ** P) s" + apply (tactic {* dtac (mk_sep_select_rule @{context} true (3,1)) 1 *}) + apply (tactic {* rtac (mk_sep_select_rule @{context} false (4,2)) 1 *}) + (* now sep_conj_impl1 can be used *) + apply (erule (1) sep_conj_impl) + oops + + +text {* Conjunct selection *} + +lemma "(A ** B ** Q ** P) s" + apply (sep_select 1) + apply (sep_select 3) + apply (sep_select 4) + oops + +lemma "\ also unrelated; (A ** B ** Q ** P) s \ \ unrelated" + apply (sep_select_asm 2) + oops + + +section {* Test cases for @{text sep_cancel}. *} + +lemma + assumes forward: "\s g p v. A g p v s \ AA g p s " + shows "\xv yv P s y x s. (A g x yv ** A g y yv ** P) s \ (AA g y ** sep_true) s" + by (sep_cancel add: forward) + +lemma + assumes forward: "\s. generic s \ instance s" + shows "(A ** generic ** B) s \ (instance ** sep_true) s" + by (sep_cancel add: forward) + +lemma "\ (A ** B) sa ; (A ** Y) s \ \ (A ** X) s" + apply (sep_cancel) + oops + +lemma "\ (A ** B) sa ; (A ** Y) s \ \ (\s. (A ** X) s) s" + apply (sep_cancel) + oops + +schematic_lemma "\ (B ** A ** C) s \ \ (\s. (A ** ?X) s) s" + by (sep_cancel) + +(* test backtracking on premises with same state *) +lemma + assumes forward: "\s. generic s \ instance s" + shows "\ (A ** B) s ; (generic ** Y) s \ \ (X ** instance) s" + apply (sep_cancel add: forward) + oops + +lemma + assumes forward: "\s. generic s \ instance s" + shows "generic s \ instance s" + by (sep_cancel add: forward) + +lemma + assumes forward: "\s. generic s \ instance s" + assumes forward2: "\s. instance s \ instance2 s" + shows "generic s \ (instance2 ** sep_true) s" + by (sep_cancel_blast add: forward forward2) + +end + diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/ex/Simple_Separation_Example.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/ex/Simple_Separation_Example.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,109 @@ +(* Title: Adaptation of example from HOL/Hoare/Separation + Author: Gerwin Klein, 2012 + Maintainers: Gerwin Klein + Rafal Kolanski +*) + +header "Example from HOL/Hoare/Separation" + +theory Simple_Separation_Example + imports "~~/src/HOL/Hoare/Hoare_Logic_Abort" "../Sep_Heap_Instance" + "../Sep_Tactics" +begin + +declare [[syntax_ambiguity_warning = false]] + +type_synonym heap = "(nat \ nat option)" + +(* This syntax isn't ideal, but this is what is used in the original *) +definition maps_to:: "nat \ nat \ heap \ bool" ("_ \ _" [56,51] 56) + where "x \ y \ \h. h = [x \ y]" + +(* If you don't mind syntax ambiguity: *) +notation pred_ex (binder "\" 10) + +(* could be generated automatically *) +definition maps_to_ex :: "nat \ heap \ bool" ("_ \ -" [56] 56) + where "x \ - \ \y. x \ y" + + +(* could be generated automatically *) +lemma maps_to_maps_to_ex [elim!]: + "(p \ v) s \ (p \ -) s" + by (auto simp: maps_to_ex_def) + +(* The basic properties of maps_to: *) +lemma maps_to_write: + "(p \ - ** P) H \ (p \ v ** P) (H (p \ v))" + apply (clarsimp simp: sep_conj_def maps_to_def maps_to_ex_def split: option.splits) + apply (rule_tac x=y in exI) + apply (auto simp: sep_disj_fun_def map_convs map_add_def split: option.splits) + done + +lemma points_to: + "(p \ v ** P) H \ the (H p) = v" + by (auto elim!: sep_conjE + simp: sep_disj_fun_def maps_to_def map_convs map_add_def + split: option.splits) + + +(* This differs from the original and uses separation logic for the definition. *) +primrec + list :: "nat \ nat list \ heap \ bool" +where + "list i [] = (\i=0\ and \)" +| "list i (x#xs) = (\i=x \ i\0\ and (EXS j. i \ j ** list j xs))" + +lemma list_empty [simp]: + shows "list 0 xs = (\s. xs = [] \ \ s)" + by (cases xs) auto + +(* The examples from Hoare/Separation.thy *) +lemma "VARS x y z w h + {(x \ y ** z \ w) h} + SKIP + {x \ z}" + apply vcg + apply(auto elim!: sep_conjE simp: maps_to_def sep_disj_fun_def domain_conv) +done + +lemma "VARS H x y z w + {(P ** Q) H} + SKIP + {(Q ** P) H}" + apply vcg + apply(simp add: sep_conj_commute) +done + +lemma "VARS H + {p\0 \ (p \ - ** list q qs) H} + H := H(p \ q) + {list p (p#qs) H}" + apply vcg + apply (auto intro: maps_to_write) +done + +lemma "VARS H p q r + {(list p Ps ** list q Qs) H} + WHILE p \ 0 + INV {\ps qs. (list p ps ** list q qs) H \ rev ps @ qs = rev Ps @ Qs} + DO r := p; p := the(H p); H := H(r \ q); q := r OD + {list q (rev Ps @ Qs) H}" + apply vcg + apply fastforce + apply clarsimp + apply (case_tac ps, simp) + apply (rename_tac p ps') + apply (clarsimp simp: sep_conj_exists sep_conj_ac) + apply (sep_subst points_to) + apply (rule_tac x = "ps'" in exI) + apply (rule_tac x = "p # qs" in exI) + apply (simp add: sep_conj_exists sep_conj_ac) + apply (rule exI) + apply (sep_rule maps_to_write) + apply (sep_cancel)+ + apply ((sep_cancel add: maps_to_maps_to_ex)+)[1] + apply clarsimp + done + +end diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/ex/Simple_Separation_Example.thy~ --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/ex/Simple_Separation_Example.thy~ Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,109 @@ +(* Title: Adaptation of example from HOL/Hoare/Separation + Author: Gerwin Klein, 2012 + Maintainers: Gerwin Klein + Rafal Kolanski +*) + +header "Example from HOL/Hoare/Separation" + +theory Simple_Separation_Example + imports "~~/src/HOL/Hoare/Hoare_Logic_Abort" "../Sep_Heap_Instance" + "../Sep_Tactics" +begin + +declare [[syntax_ambiguity_warning = false]] + +type_synonym heap = "(nat \ nat option)" + +(* This syntax isn't ideal, but this is what is used in the original *) +definition maps_to:: "nat \ nat \ heap \ bool" ("_ \ _" [56,51] 56) + where "x \ y \ \h. h = [x \ y]" + +(* If you don't mind syntax ambiguity: *) +notation pred_ex (binder "\" 10) + +(* could be generated automatically *) +definition maps_to_ex :: "nat \ heap \ bool" ("_ \ -" [56] 56) + where "x \ - \ \y. x \ y" + + +(* could be generated automatically *) +lemma maps_to_maps_to_ex [elim!]: + "(p \ v) s \ (p \ -) s" + by (auto simp: maps_to_ex_def) + +(* The basic properties of maps_to: *) +lemma maps_to_write: + "(p \ - ** P) H \ (p \ v ** P) (H (p \ v))" + apply (clarsimp simp: sep_conj_def maps_to_def maps_to_ex_def split: option.splits) + apply (rule_tac x=y in exI) + apply (auto simp: sep_disj_fun_def map_convs map_add_def split: option.splits) + done + +lemma points_to: + "(p \ v ** P) H \ the (H p) = v" + by (auto elim!: sep_conjE + simp: sep_disj_fun_def maps_to_def map_convs map_add_def + split: option.splits) + + +(* This differs from the original and uses separation logic for the definition. *) +primrec + list :: "nat \ nat list \ heap \ bool" +where + "list i [] = (\i=0\ and \)" +| "list i (x#xs) = (\i=x \ i\0\ and (EXS j. i \ j ** list j xs))" + +lemma list_empty [simp]: + shows "list 0 xs = (\s. xs = [] \ \ s)" + by (cases xs) auto + +(* The examples from Hoare/Separation.thy *) +lemma "VARS x y z w h + {(x \ y ** z \ w) h} + SKIP + {x \ z}" + apply vcg + apply(auto elim!: sep_conjE simp: maps_to_def sep_disj_fun_def domain_conv) +done + +lemma "VARS H x y z w + {(P ** Q) H} + SKIP + {(Q ** P) H}" + apply vcg + apply(simp add: sep_conj_commute) +done + +lemma "VARS H + {p\0 \ (p \ - ** list q qs) H} + H := H(p \ q) + {list p (p#qs) H}" + apply vcg + apply (auto intro: maps_to_write) +done + +lemma "VARS H p q r + {(list p Ps ** list q Qs) H} + WHILE p \ 0 + INV {\ps qs. (list p ps ** list q qs) H \ rev ps @ qs = rev Ps @ Qs} + DO r := p; p := the(H p); H := H(r \ q); q := r OD + {list q (rev Ps @ Qs) H}" + apply vcg + apply fastforce + apply clarsimp + apply (case_tac ps, simp) + apply (rename_tac p ps') + apply (clarsimp simp: sep_conj_exists sep_conj_ac) + apply (sep_subst points_to) + apply (rule_tac x = "ps'" in exI) + apply (rule_tac x = "p # qs" in exI) + apply (simp add: sep_conj_exists sep_conj_ac) + apply (rule exI) + apply (sep_rule maps_to_write) + apply (sep_cancel)+ + apply ((sep_cancel add: maps_to_maps_to_ex)+)[1] + apply clarsimp + done + +end diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/ex/VM_Example.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/ex/VM_Example.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,78 @@ +(* Title: Adaptation of example from HOL/Hoare/Separation + Author: Rafal Kolanski, 2012 + Maintainers: Gerwin Klein + Rafal Kolanski +*) + +header "Separation Algebra for Virtual Memory" + +theory VM_Example +imports "../Sep_Tactics" "../Map_Extra" +begin + +text {* + Example instantiation of the abstract separation algebra to the sliced-memory + model used for building a separation logic in ``Verification of Programs in + Virtual Memory Using Separation Logic'' (PhD Thesis) by Rafal Kolanski. + + We wrap up the concept of physical and virtual pointers as well as value + (usually a byte), and the page table root, into a datatype for instantiation. + This avoids having to produce a hierarchy of type classes. + + The result is more general than the original. It does not mention the types + of pointers or virtual memory addresses. Instead of supporting only + singleton page table roots, we now support sets so we can identify a single + 0 for the monoid. + This models multiple page tables in memory, whereas the original logic was + only capable of one at a time. +*} + +datatype ('p,'v,'value,'r) vm_sep_state + = VMSepState "((('p \ 'v) \ 'value) \ 'r set)" + +instantiation vm_sep_state :: (type, type, type, type) sep_algebra +begin + +fun + vm_heap :: "('a,'b,'c,'d) vm_sep_state \ (('a \ 'b) \ 'c)" where + "vm_heap (VMSepState (h,r)) = h" + +fun + vm_root :: "('a,'b,'c,'d) vm_sep_state \ 'd set" where + "vm_root (VMSepState (h,r)) = r" + +definition + sep_disj_vm_sep_state :: "('a, 'b, 'c, 'd) vm_sep_state + \ ('a, 'b, 'c, 'd) vm_sep_state \ bool" where + "sep_disj_vm_sep_state x y = vm_heap x \ vm_heap y" + +definition + zero_vm_sep_state :: "('a, 'b, 'c, 'd) vm_sep_state" where + "zero_vm_sep_state \ VMSepState (empty, {})" + +fun + plus_vm_sep_state :: "('a, 'b, 'c, 'd) vm_sep_state + \ ('a, 'b, 'c, 'd) vm_sep_state + \ ('a, 'b, 'c, 'd) vm_sep_state" where + "plus_vm_sep_state (VMSepState (x,r)) (VMSepState (y,r')) + = VMSepState (x ++ y, r \ r')" + +instance + apply default + apply (simp add: zero_vm_sep_state_def sep_disj_vm_sep_state_def) + apply (fastforce simp: sep_disj_vm_sep_state_def map_disj_def) + apply (case_tac x, clarsimp simp: zero_vm_sep_state_def) + apply (case_tac x, case_tac y) + apply (fastforce simp: sep_disj_vm_sep_state_def map_add_ac) + apply (case_tac x, case_tac y, case_tac z) + apply (fastforce simp: sep_disj_vm_sep_state_def) + apply (case_tac x, case_tac y, case_tac z) + apply (fastforce simp: sep_disj_vm_sep_state_def map_add_disj) + apply (case_tac x, case_tac y, case_tac z) + apply (fastforce simp: sep_disj_vm_sep_state_def map_add_disj map_disj_com) + done + +end + +end + diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/ex/capDL/Abstract_Separation_D.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/ex/capDL/Abstract_Separation_D.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,714 @@ +(* Author: Andrew Boyton, 2012 + Maintainers: Gerwin Klein + Rafal Kolanski +*) + +header "Instantiating capDL as a separation algebra." + +theory Abstract_Separation_D +imports "../../Sep_Tactics" Types_D "../../Map_Extra" +begin + +(************************************** + * Start of lemmas to move elsewhere. * + **************************************) + +lemma inter_empty_not_both: +"\x \ A; A \ B = {}\ \ x \ B" + by fastforce + +lemma union_intersection: + "A \ (A \ B) = A" + "B \ (A \ B) = B" + "(A \ B) \ A = A" + "(A \ B) \ B = B" + by fastforce+ + +lemma union_intersection1: "A \ (A \ B) = A" + by (rule inf_sup_absorb) +lemma union_intersection2: "B \ (A \ B) = B" + by fastforce + +(* This lemma is strictly weaker than restrict_map_disj. *) +lemma restrict_map_disj': + "S \ T = {} \ h |` S \ h' |` T" + by (auto simp: map_disj_def restrict_map_def dom_def) + +lemma map_add_restrict_comm: + "S \ T = {} \ h |` S ++ h' |` T = h' |` T ++ h |` S" + apply (drule restrict_map_disj') + apply (erule map_add_com) + done + +(************************************ + * End of lemmas to move elsewhere. * + ************************************) + + + +(* The state for separation logic has: + * The memory heap. + * A function for which objects own which fields. + In capDL, we say that an object either owns all of its fields, or none of them. + These are both taken from the cdl_state. + *) + +datatype sep_state = SepState cdl_heap cdl_ghost_state + +(* Functions to get the heap and the ghost_state from the sep_state. *) +primrec sep_heap :: "sep_state \ cdl_heap" +where "sep_heap (SepState h gs) = h" + +primrec sep_ghost_state :: "sep_state \ cdl_ghost_state" +where "sep_ghost_state (SepState h gs) = gs" + +definition + the_set :: "'a option set \ 'a set" +where + "the_set xs = {x. Some x \ xs}" + +lemma the_set_union [simp]: + "the_set (A \ B) = the_set A \ the_set B" + by (fastforce simp: the_set_def) + +lemma the_set_inter [simp]: + "the_set (A \ B) = the_set A \ the_set B" + by (fastforce simp: the_set_def) + +lemma the_set_inter_empty: + "A \ B = {} \ the_set A \ the_set B = {}" + by (fastforce simp: the_set_def) + + +(* As the capDL operations mostly take the state (rather than the heap) + * we need to redefine some of them again to take just the heap. + *) +definition + slots_of_heap :: "cdl_heap \ cdl_object_id \ cdl_cap_map" +where + "slots_of_heap h \ \obj_id. + case h obj_id of + None \ empty + | Some obj \ object_slots obj" + +(* Adds new caps to an object. It won't overwrite on a collision. *) +definition + add_to_slots :: "cdl_cap_map \ cdl_object \ cdl_object" +where + "add_to_slots new_val obj \ update_slots (new_val ++ (object_slots obj)) obj" + +lemma add_to_slots_assoc: + "add_to_slots x (add_to_slots (y ++ z) obj) = + add_to_slots (x ++ y) (add_to_slots z obj)" + apply (clarsimp simp: add_to_slots_def update_slots_def object_slots_def) + apply (fastforce simp: cdl_tcb.splits cdl_cnode.splits + split: cdl_object.splits) + done + +(* Lemmas about add_to_slots, update_slots and object_slots. *) +lemma add_to_slots_twice [simp]: + "add_to_slots x (add_to_slots y a) = add_to_slots (x ++ y) a" + by (fastforce simp: add_to_slots_def update_slots_def object_slots_def + split: cdl_object.splits) + +lemma slots_of_heap_empty [simp]: "slots_of_heap empty object_id = empty" + by (simp add: slots_of_heap_def) + +lemma slots_of_heap_empty2 [simp]: + "h obj_id = None \ slots_of_heap h obj_id = empty" + by (simp add: slots_of_heap_def) + +lemma update_slots_add_to_slots_empty [simp]: + "update_slots empty (add_to_slots new obj) = update_slots empty obj" + by (clarsimp simp: update_slots_def add_to_slots_def split:cdl_object.splits) + +lemma update_object_slots_id [simp]: "update_slots (object_slots a) a = a" + by (clarsimp simp: update_slots_def object_slots_def + split: cdl_object.splits) + +lemma update_slots_of_heap_id [simp]: + "h obj_id = Some obj \ update_slots (slots_of_heap h obj_id) obj = obj" + by (clarsimp simp: update_slots_def slots_of_heap_def object_slots_def + split: cdl_object.splits) + +lemma add_to_slots_empty [simp]: "add_to_slots empty h = h" + by (simp add: add_to_slots_def) + +lemma update_slots_eq: + "update_slots a o1 = update_slots a o2 \ update_slots b o1 = update_slots b o2" + by (fastforce simp: update_slots_def cdl_tcb.splits cdl_cnode.splits + split: cdl_object.splits) + + + +(* If there are not two conflicting objects at a position in two states. + * Objects conflict if their types are different or their ghost_states collide. + *) +definition + not_conflicting_objects :: "sep_state \ sep_state \ cdl_object_id \ bool" +where + "not_conflicting_objects state_a state_b = (\obj_id. + let heap_a = sep_heap state_a; + heap_b = sep_heap state_b; + gs_a = sep_ghost_state state_a; + gs_b = sep_ghost_state state_b + in case (heap_a obj_id, heap_b obj_id) of + (Some o1, Some o2) \ object_type o1 = object_type o2 \ gs_a obj_id \ gs_b obj_id = {} + | _ \ True)" + + +(* "Cleans" slots to conform with the compontents. *) +definition + clean_slots :: "cdl_cap_map \ cdl_components \ cdl_cap_map" +where + "clean_slots slots cmp \ slots |` the_set cmp" + +(* Sets the fields of an object to a "clean" state. + Because a frame's size is part of it's type, we don't reset it. *) +definition + object_clean_fields :: "cdl_object \ cdl_components \ cdl_object" +where + "object_clean_fields obj cmp \ if None \ cmp then obj else case obj of + Tcb x \ Tcb (x\cdl_tcb_fault_endpoint := undefined\) + | CNode x \ CNode (x\cdl_cnode_size_bits := undefined \) + | _ \ obj" + +(* Sets the slots of an object to a "clean" state. *) +definition + object_clean_slots :: "cdl_object \ cdl_components \ cdl_object" +where + "object_clean_slots obj cmp \ update_slots (clean_slots (object_slots obj) cmp) obj" + +(* Sets an object to a "clean" state. *) +definition + object_clean :: "cdl_object \ cdl_components \ cdl_object" +where + "object_clean obj gs \ object_clean_slots (object_clean_fields obj gs) gs" + +(* Overrides the left object with the attributes of the right, as specified by the ghost state. + If the components for an object are empty, then this object is treated as empty, and thus ignored. + *) +definition + object_add :: "cdl_object \ cdl_object \ cdl_components \ cdl_components \ cdl_object" +where + "object_add obj_a obj_b cmps_a cmps_b \ + let clean_obj_a = object_clean obj_a cmps_a; + clean_obj_b = object_clean obj_b cmps_b + in if (cmps_a = {}) + then clean_obj_b + else if (cmps_b = {}) + then clean_obj_a + else if (None \ cmps_b) + then (update_slots (object_slots clean_obj_a ++ object_slots clean_obj_b) clean_obj_b) + else (update_slots (object_slots clean_obj_a ++ object_slots clean_obj_b) clean_obj_a)" + +(* Heaps are added by adding their repsective objects. + * The ghost state tells us which object's fields should be taken. + * Adding objects of the same type adds their caps + * (overwrites the left with the right). + *) +definition + cdl_heap_add :: "sep_state \ sep_state \ cdl_heap" +where + "cdl_heap_add state_a state_b \ \obj_id. + let heap_a = sep_heap state_a; + heap_b = sep_heap state_b; + gs_a = sep_ghost_state state_a; + gs_b = sep_ghost_state state_b + in case heap_b obj_id of + None \ heap_a obj_id + | Some obj_b \ case heap_a obj_id of + None \ heap_b obj_id + | Some obj_a \ Some (object_add obj_a obj_b (gs_a obj_id) (gs_b obj_id))" + +(* Heaps are added by adding their repsective objects. + * The ghost state tells us which object's fields should be taken. + * Adding objects of the same type adds their caps + * (overwrites the left with the right). + *) +definition + cdl_ghost_state_add :: "sep_state \ sep_state \ cdl_ghost_state" +where + "cdl_ghost_state_add state_a state_b \ \obj_id. + let heap_a = sep_heap state_a; + heap_b = sep_heap state_b; + gs_a = sep_ghost_state state_a; + gs_b = sep_ghost_state state_b + in if heap_a obj_id = None \ heap_b obj_id \ None then gs_b obj_id + else if heap_b obj_id = None \ heap_a obj_id \ None then gs_a obj_id + else gs_a obj_id \ gs_b obj_id" + + +(* Adding states adds their heaps, + * and each objects owns whichever fields it owned in either heap. + *) +definition + sep_state_add :: "sep_state \ sep_state \ sep_state" +where + "sep_state_add state_a state_b \ + let + heap_a = sep_heap state_a; + heap_b = sep_heap state_b; + gs_a = sep_ghost_state state_a; + gs_b = sep_ghost_state state_b + in + SepState (cdl_heap_add state_a state_b) (cdl_ghost_state_add state_a state_b)" + + +(* Heaps are disjoint if for all of their objects: + * the caps of their respective objects are disjoint, + * their respective objects don't conflict, + * they don't both own any of the same fields. +*) +definition + sep_state_disj :: "sep_state \ sep_state \ bool" +where + "sep_state_disj state_a state_b \ + let + heap_a = sep_heap state_a; + heap_b = sep_heap state_b; + gs_a = sep_ghost_state state_a; + gs_b = sep_ghost_state state_b + in + \obj_id. not_conflicting_objects state_a state_b obj_id" + +lemma not_conflicting_objects_comm: + "not_conflicting_objects h1 h2 obj = not_conflicting_objects h2 h1 obj" + apply (clarsimp simp: not_conflicting_objects_def split:option.splits) + apply (fastforce simp: update_slots_def cdl_tcb.splits cdl_cnode.splits + split: cdl_object.splits) + done + +lemma object_clean_comm: + "\object_type obj_a = object_type obj_b; + object_slots obj_a ++ object_slots obj_b = object_slots obj_b ++ object_slots obj_a; None \ cmp\ + \ object_clean (add_to_slots (object_slots obj_a) obj_b) cmp = + object_clean (add_to_slots (object_slots obj_b) obj_a) cmp" + apply (clarsimp simp: object_type_def split: cdl_object.splits) + apply (clarsimp simp: object_clean_def object_clean_slots_def object_clean_fields_def + add_to_slots_def object_slots_def update_slots_def + cdl_tcb.splits cdl_cnode.splits + split: cdl_object.splits)+ + done + +lemma add_to_slots_object_slots: + "object_type y = object_type z + \ add_to_slots (object_slots (add_to_slots (x) y)) z = + add_to_slots (x ++ object_slots y) z" + apply (clarsimp simp: add_to_slots_def update_slots_def object_slots_def) + apply (fastforce simp: object_type_def cdl_tcb.splits cdl_cnode.splits + split: cdl_object.splits) + done + +lemma not_conflicting_objects_empty [simp]: + "not_conflicting_objects s (SepState empty (\obj_id. {})) obj_id" + by (clarsimp simp: not_conflicting_objects_def split:option.splits) + +lemma empty_not_conflicting_objects [simp]: + "not_conflicting_objects (SepState empty (\obj_id. {})) s obj_id" + by (clarsimp simp: not_conflicting_objects_def split:option.splits) + +lemma not_conflicting_objects_empty_object [elim!]: + "(sep_heap x) obj_id = None \ not_conflicting_objects x y obj_id" + by (clarsimp simp: not_conflicting_objects_def) + +lemma empty_object_not_conflicting_objects [elim!]: + "(sep_heap y) obj_id = None \ not_conflicting_objects x y obj_id" + apply (drule not_conflicting_objects_empty_object [where y=x]) + apply (clarsimp simp: not_conflicting_objects_comm) + done + +lemma cdl_heap_add_empty [simp]: + "cdl_heap_add (SepState h gs) (SepState empty (\obj_id. {})) = h" + by (simp add: cdl_heap_add_def) + +lemma empty_cdl_heap_add [simp]: + "cdl_heap_add (SepState empty (\obj_id. {})) (SepState h gs)= h" + apply (simp add: cdl_heap_add_def) + apply (rule ext) + apply (clarsimp split: option.splits) + done + +lemma map_add_result_empty1: "a ++ b = empty \ a = empty" + apply (subgoal_tac "dom (a++b) = {}") + apply (subgoal_tac "dom (a) = {}") + apply clarsimp + apply (unfold dom_map_add)[1] + apply clarsimp + apply clarsimp + done + +lemma map_add_result_empty2: "a ++ b = empty \ b = empty" + apply (subgoal_tac "dom (a++b) = {}") + apply (subgoal_tac "dom (a) = {}") + apply clarsimp + apply (unfold dom_map_add)[1] + apply clarsimp + apply clarsimp + done + +lemma map_add_emptyE [elim!]: "\a ++ b = empty; \a = empty; b = empty\ \ R\ \ R" + apply (frule map_add_result_empty1) + apply (frule map_add_result_empty2) + apply clarsimp + done + +lemma clean_slots_empty [simp]: + "clean_slots empty cmp = empty" + by (clarsimp simp: clean_slots_def) + +lemma object_type_update_slots [simp]: + "object_type (update_slots slots x) = object_type x" + by (clarsimp simp: object_type_def update_slots_def split: cdl_object.splits) + +lemma object_type_object_clean_slots [simp]: + "object_type (object_clean_slots x cmp) = object_type x" + by (clarsimp simp: object_clean_slots_def) + +lemma object_type_object_clean_fields [simp]: + "object_type (object_clean_fields x cmp) = object_type x" + by (clarsimp simp: object_clean_fields_def object_type_def split: cdl_object.splits) + +lemma object_type_object_clean [simp]: + "object_type (object_clean x cmp) = object_type x" + by (clarsimp simp: object_clean_def) + +lemma object_type_add_to_slots [simp]: + "object_type (add_to_slots slots x) = object_type x" + by (clarsimp simp: object_type_def add_to_slots_def update_slots_def split: cdl_object.splits) + +lemma object_slots_update_slots [simp]: + "has_slots obj \ object_slots (update_slots slots obj) = slots" + by (clarsimp simp: object_slots_def update_slots_def has_slots_def + split: cdl_object.splits) + +lemma object_slots_update_slots_empty [simp]: + "\has_slots obj \ object_slots (update_slots slots obj) = empty" + by (clarsimp simp: object_slots_def update_slots_def has_slots_def + split: cdl_object.splits) + +lemma update_slots_no_slots [simp]: + "\has_slots obj \ update_slots slots obj = obj" + by (clarsimp simp: update_slots_def has_slots_def split: cdl_object.splits) + +lemma update_slots_update_slots [simp]: + "update_slots slots (update_slots slots' obj) = update_slots slots obj" + by (clarsimp simp: update_slots_def split: cdl_object.splits) + +lemma update_slots_same_object: + "a = b \ update_slots a obj = update_slots b obj" + by (erule arg_cong) + +lemma object_type_has_slots: + "\has_slots x; object_type x = object_type y\ \ has_slots y" + by (clarsimp simp: object_type_def has_slots_def split: cdl_object.splits) + +lemma object_slots_object_clean_fields [simp]: + "object_slots (object_clean_fields obj cmp) = object_slots obj" + by (clarsimp simp: object_slots_def object_clean_fields_def split: cdl_object.splits) + +lemma object_slots_object_clean_slots [simp]: + "object_slots (object_clean_slots obj cmp) = clean_slots (object_slots obj) cmp" + by (clarsimp simp: object_clean_slots_def object_slots_def update_slots_def split: cdl_object.splits) + +lemma object_slots_object_clean [simp]: + "object_slots (object_clean obj cmp) = clean_slots (object_slots obj) cmp" + by (clarsimp simp: object_clean_def) + +lemma object_slots_add_to_slots [simp]: + "object_type y = object_type z \ object_slots (add_to_slots (object_slots y) z) = object_slots y ++ object_slots z" + by (clarsimp simp: object_slots_def add_to_slots_def update_slots_def object_type_def split: cdl_object.splits) + +lemma update_slots_object_clean_slots [simp]: + "update_slots slots (object_clean_slots obj cmp) = update_slots slots obj" + by (clarsimp simp: object_clean_slots_def) + +lemma object_clean_fields_idem [simp]: + "object_clean_fields (object_clean_fields obj cmp) cmp = object_clean_fields obj cmp" + by (clarsimp simp: object_clean_fields_def split: cdl_object.splits) + +lemma object_clean_slots_idem [simp]: + "object_clean_slots (object_clean_slots obj cmp) cmp = object_clean_slots obj cmp" + apply (case_tac "has_slots obj") + apply (clarsimp simp: object_clean_slots_def clean_slots_def)+ + done + +lemma object_clean_fields_object_clean_slots [simp]: + "object_clean_fields (object_clean_slots obj gs) gs = object_clean_slots (object_clean_fields obj gs) gs" + by (clarsimp simp: object_clean_fields_def object_clean_slots_def + clean_slots_def object_slots_def update_slots_def + split: cdl_object.splits) + +lemma object_clean_idem [simp]: + "object_clean (object_clean obj cmp) cmp = object_clean obj cmp" + by (clarsimp simp: object_clean_def) + +lemma has_slots_object_clean_slots: + "has_slots (object_clean_slots obj cmp) = has_slots obj" + by (clarsimp simp: has_slots_def object_clean_slots_def update_slots_def split: cdl_object.splits) + +lemma has_slots_object_clean_fields: + "has_slots (object_clean_fields obj cmp) = has_slots obj" + by (clarsimp simp: has_slots_def object_clean_fields_def split: cdl_object.splits) + +lemma has_slots_object_clean: + "has_slots (object_clean obj cmp) = has_slots obj" + by (clarsimp simp: object_clean_def has_slots_object_clean_slots has_slots_object_clean_fields) + +lemma object_slots_update_slots_object_clean_fields [simp]: + "object_slots (update_slots slots (object_clean_fields obj cmp)) = object_slots (update_slots slots obj)" + apply (case_tac "has_slots obj") + apply (clarsimp simp: has_slots_object_clean_fields)+ + done + +lemma object_clean_fields_update_slots [simp]: + "object_clean_fields (update_slots slots obj) cmp = update_slots slots (object_clean_fields obj cmp)" + by (clarsimp simp: object_clean_fields_def update_slots_def split: cdl_object.splits) + +lemma object_clean_fields_twice [simp]: + "(object_clean_fields (object_clean_fields obj cmp') cmp) = object_clean_fields obj (cmp \ cmp')" + by (clarsimp simp: object_clean_fields_def split: cdl_object.splits) + +lemma update_slots_object_clean_fields: + "\None \ cmps; None \ cmps'; object_type obj = object_type obj'\ + \ update_slots slots (object_clean_fields obj cmps) = + update_slots slots (object_clean_fields obj' cmps')" + by (fastforce simp: update_slots_def object_clean_fields_def object_type_def split: cdl_object.splits) + +lemma object_clean_fields_no_slots: + "\None \ cmps; None \ cmps'; object_type obj = object_type obj'; \ has_slots obj; \ has_slots obj'\ + \ object_clean_fields obj cmps = object_clean_fields obj' cmps'" + by (fastforce simp: object_clean_fields_def object_type_def has_slots_def split: cdl_object.splits) + +lemma update_slots_object_clean: + "\None \ cmps; None \ cmps'; object_type obj = object_type obj'\ + \ update_slots slots (object_clean obj cmps) = update_slots slots (object_clean obj' cmps')" + apply (clarsimp simp: object_clean_def object_clean_slots_def) + apply (erule (2) update_slots_object_clean_fields) + done + +lemma cdl_heap_add_assoc': + "\obj_id. not_conflicting_objects x z obj_id \ + not_conflicting_objects y z obj_id \ + not_conflicting_objects x z obj_id \ + cdl_heap_add (SepState (cdl_heap_add x y) (cdl_ghost_state_add x y)) z = + cdl_heap_add x (SepState (cdl_heap_add y z) (cdl_ghost_state_add y z))" + apply (rule ext) + apply (rename_tac obj_id) + apply (erule_tac x=obj_id in allE) + apply (clarsimp simp: cdl_heap_add_def cdl_ghost_state_add_def not_conflicting_objects_def) + apply (simp add: Let_unfold split: option.splits) + apply (rename_tac obj_y obj_x obj_z) + apply (clarsimp simp: object_add_def clean_slots_def object_clean_def object_clean_slots_def Let_unfold) + apply (case_tac "has_slots obj_z") + apply (subgoal_tac "has_slots obj_y") + apply (subgoal_tac "has_slots obj_x") + apply ((clarsimp simp: has_slots_object_clean_fields has_slots_object_clean_slots has_slots_object_clean + map_add_restrict union_intersection | + drule inter_empty_not_both | + erule update_slots_object_clean_fields | + erule object_type_has_slots, simp | + simp | safe)+)[3] + apply (subgoal_tac "\ has_slots obj_y") + apply (subgoal_tac "\ has_slots obj_x") + apply ((clarsimp simp: has_slots_object_clean_fields has_slots_object_clean_slots has_slots_object_clean + map_add_restrict union_intersection | + drule inter_empty_not_both | + erule object_clean_fields_no_slots | + erule object_type_has_slots, simp | + simp | safe)+) + apply (fastforce simp: object_type_has_slots)+ + done + +lemma cdl_heap_add_assoc: + "\sep_state_disj x y; sep_state_disj y z; sep_state_disj x z\ + \ cdl_heap_add (SepState (cdl_heap_add x y) (cdl_ghost_state_add x y)) z = + cdl_heap_add x (SepState (cdl_heap_add y z) (cdl_ghost_state_add y z))" + apply (clarsimp simp: sep_state_disj_def) + apply (cut_tac cdl_heap_add_assoc') + apply fast + apply fastforce + done + +lemma cdl_ghost_state_add_assoc: + "cdl_ghost_state_add (SepState (cdl_heap_add x y) (cdl_ghost_state_add x y)) z = + cdl_ghost_state_add x (SepState (cdl_heap_add y z) (cdl_ghost_state_add y z))" + apply (rule ext) + apply (fastforce simp: cdl_heap_add_def cdl_ghost_state_add_def Let_unfold) + done + +lemma clean_slots_map_add_comm: + "cmps_a \ cmps_b = {} + \ clean_slots slots_a cmps_a ++ clean_slots slots_b cmps_b = + clean_slots slots_b cmps_b ++ clean_slots slots_a cmps_a" + apply (clarsimp simp: clean_slots_def) + apply (drule the_set_inter_empty) + apply (erule map_add_restrict_comm) + done + +lemma object_clean_all: + "object_type obj_a = object_type obj_b \ object_clean obj_b {} = object_clean obj_a {}" + apply (clarsimp simp: object_clean_def object_clean_slots_def clean_slots_def the_set_def) + apply (rule_tac cmps'1="{}" and obj'1="obj_a" in trans [OF update_slots_object_clean_fields], fastforce+) + done + +lemma object_add_comm: + "\object_type obj_a = object_type obj_b; cmps_a \ cmps_b = {}\ + \ object_add obj_a obj_b cmps_a cmps_b = object_add obj_b obj_a cmps_b cmps_a" + apply (clarsimp simp: object_add_def Let_unfold) + apply (rule conjI | clarsimp)+ + apply fastforce + apply (rule conjI | clarsimp)+ + apply (drule_tac slots_a = "object_slots obj_a" and slots_b = "object_slots obj_b" in clean_slots_map_add_comm) + apply fastforce + apply (rule conjI | clarsimp)+ + apply (drule_tac slots_a = "object_slots obj_a" and slots_b = "object_slots obj_b" in clean_slots_map_add_comm) + apply fastforce + apply (rule conjI | clarsimp)+ + apply (erule object_clean_all) + apply (clarsimp) + apply (rule_tac cmps'1=cmps_b and obj'1=obj_b in trans [OF update_slots_object_clean], assumption+) + apply (drule_tac slots_a = "object_slots obj_a" and slots_b = "object_slots obj_b" in clean_slots_map_add_comm) + apply fastforce + done + +lemma sep_state_add_comm: + "sep_state_disj x y \ sep_state_add x y = sep_state_add y x" + apply (clarsimp simp: sep_state_add_def sep_state_disj_def) + apply (rule conjI) + apply (case_tac x, case_tac y, clarsimp) + apply (rename_tac heap_a gs_a heap_b gs_b) + apply (clarsimp simp: cdl_heap_add_def Let_unfold) + apply (rule ext) + apply (case_tac "heap_a obj_id") + apply (case_tac "heap_b obj_id", simp_all add: slots_of_heap_def) + apply (case_tac "heap_b obj_id", simp_all add: slots_of_heap_def) + apply (rename_tac obj_a obj_b) + apply (erule_tac x=obj_id in allE) + apply (rule object_add_comm) + apply (clarsimp simp: not_conflicting_objects_def) + apply (clarsimp simp: not_conflicting_objects_def) + apply (rule ext, fastforce simp: cdl_ghost_state_add_def Let_unfold Un_commute) + done + +lemma add_to_slots_comm: + "\object_slots y_obj \ object_slots z_obj; update_slots empty y_obj = update_slots empty z_obj \ + \ add_to_slots (object_slots z_obj) y_obj = add_to_slots (object_slots y_obj) z_obj" + by (fastforce simp: add_to_slots_def update_slots_def object_slots_def + cdl_tcb.splits cdl_cnode.splits + dest!: map_add_com + split: cdl_object.splits) + +lemma cdl_heap_add_none1: + "cdl_heap_add x y obj_id = None \ (sep_heap x) obj_id = None" + by (clarsimp simp: cdl_heap_add_def Let_unfold split:option.splits split_if_asm) + +lemma cdl_heap_add_none2: + "cdl_heap_add x y obj_id = None \ (sep_heap y) obj_id = None" + by (clarsimp simp: cdl_heap_add_def Let_unfold split:option.splits split_if_asm) + +lemma object_type_object_addL: + "object_type obj = object_type obj' + \ object_type (object_add obj obj' cmp cmp') = object_type obj" + by (clarsimp simp: object_add_def Let_unfold) + +lemma object_type_object_addR: + "object_type obj = object_type obj' + \ object_type (object_add obj obj' cmp cmp') = object_type obj'" + by (clarsimp simp: object_add_def Let_unfold) + +lemma sep_state_add_disjL: + "\sep_state_disj y z; sep_state_disj x (sep_state_add y z)\ \ sep_state_disj x y" + apply (clarsimp simp: sep_state_disj_def sep_state_add_def) + apply (rename_tac obj_id) + apply (clarsimp simp: not_conflicting_objects_def) + apply (erule_tac x=obj_id in allE)+ + apply (fastforce simp: cdl_heap_add_def cdl_ghost_state_add_def object_type_object_addR + split: option.splits) + done + +lemma sep_state_add_disjR: + "\sep_state_disj y z; sep_state_disj x (sep_state_add y z)\ \ sep_state_disj x z" + apply (clarsimp simp: sep_state_disj_def sep_state_add_def) + apply (rename_tac obj_id) + apply (clarsimp simp: not_conflicting_objects_def) + apply (erule_tac x=obj_id in allE)+ + apply (fastforce simp: cdl_heap_add_def cdl_ghost_state_add_def object_type_object_addR + split: option.splits) + done + +lemma sep_state_add_disj: + "\sep_state_disj y z; sep_state_disj x y; sep_state_disj x z\ \ sep_state_disj x (sep_state_add y z)" + apply (clarsimp simp: sep_state_disj_def sep_state_add_def) + apply (rename_tac obj_id) + apply (clarsimp simp: not_conflicting_objects_def) + apply (erule_tac x=obj_id in allE)+ + apply (fastforce simp: cdl_heap_add_def cdl_ghost_state_add_def object_type_object_addR + split: option.splits) + done + + + + +(*********************************************) +(* Definition of separation logic for capDL. *) +(*********************************************) + +instantiation "sep_state" :: zero +begin + definition "0 \ SepState empty (\obj_id. {})" + instance .. +end + +instantiation "sep_state" :: stronger_sep_algebra +begin + +definition "(op ##) \ sep_state_disj" +definition "(op +) \ sep_state_add" + + + +(********************************************** + * The proof that this is a separation logic. * + **********************************************) + +instance + apply default +(* x ## 0 *) + apply (simp add: sep_disj_sep_state_def sep_state_disj_def zero_sep_state_def) +(* x ## y \ y ## x *) + apply (clarsimp simp: not_conflicting_objects_comm sep_disj_sep_state_def sep_state_disj_def Let_unfold + map_disj_com not_conflicting_objects_comm Int_commute) +(* x + 0 = x *) + apply (simp add: plus_sep_state_def sep_state_add_def zero_sep_state_def) + apply (case_tac x) + apply (clarsimp simp: cdl_heap_add_def) + apply (rule ext) + apply (clarsimp simp: cdl_ghost_state_add_def split:split_if_asm) +(* x ## y \ x + y = y + x *) + apply (clarsimp simp: plus_sep_state_def sep_disj_sep_state_def) + apply (erule sep_state_add_comm) +(* (x + y) + z = x + (y + z) *) + apply (simp add: plus_sep_state_def sep_state_add_def) + apply (rule conjI) + apply (clarsimp simp: sep_disj_sep_state_def) + apply (erule (2) cdl_heap_add_assoc) + apply (rule cdl_ghost_state_add_assoc) +(* x ## y + z = (x ## y \ x ## z) *) + apply (clarsimp simp: plus_sep_state_def sep_disj_sep_state_def) + apply (rule iffI) + (* x ## y + z \ (x ## y \ x ## z) *) + apply (rule conjI) + (* x ## y + z \ (x ## y) *) + apply (erule (1) sep_state_add_disjL) + (* x ## y + z \ (x ## z) *) + apply (erule (1) sep_state_add_disjR) + (* x ## y + z \ (x ## y \ x ## z) *) + apply clarsimp + apply (erule (2) sep_state_add_disj) + done + +end + + +end diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/ex/capDL/Separation_D.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/ex/capDL/Separation_D.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,72 @@ +(* Author: Andrew Boyton, 2012 + Maintainers: Gerwin Klein + Rafal Kolanski +*) + +header "Defining some separation logic maps-to predicates on top of the instantiation." + +theory Separation_D +imports Abstract_Separation_D +begin + +type_synonym sep_pred = "sep_state \ bool" + +definition + state_sep_projection :: "cdl_state \ sep_state" +where + "state_sep_projection \ \s. SepState (cdl_objects s) (cdl_ghost_state s)" + +(* This turns a separation logic predicate into a predicate on the capDL state. *) +abbreviation + lift' :: "(sep_state \ 'a) \ cdl_state \ 'a" ("<_>") +where + "

s \ P (state_sep_projection s)" + +(* The generalisation of the maps to operator for separation logic. *) +definition + sep_map_general :: "cdl_object_id \ cdl_object \ cdl_components \ sep_pred" +where + "sep_map_general p obj gs \ \s. sep_heap s = [p \ obj] \ sep_ghost_state s p = gs" + +(* Alternate definition without the [p \ obj] notation. *) +lemma sep_map_general_def2: + "sep_map_general p obj gs s = + (dom (sep_heap s) = {p} \ ko_at obj p (sep_heap s) \ sep_ghost_state s p = gs)" + apply (clarsimp simp: sep_map_general_def object_at_def) + apply (rule) + apply clarsimp + apply (clarsimp simp: fun_upd_def) + apply (rule ext) + apply (fastforce simp: dom_def split:split_if) + done + +(* There is an object there. *) +definition + sep_map_i :: "cdl_object_id \ cdl_object \ sep_pred" ("_ \i _" [76,71] 76) +where + "p \i obj \ sep_map_general p obj UNIV" + +(* The fields are there (and there are no caps). *) +definition + sep_map_f :: "cdl_object_id \ cdl_object \ sep_pred" ("_ \f _" [76,71] 76) +where + "p \f obj \ sep_map_general p (update_slots empty obj) {None}" + +(* There is that cap there. *) +definition + sep_map_c :: "cdl_cap_ref \ cdl_cap \ sep_pred" ("_ \c _" [76,71] 76) +where + "p \c cap \ \s. let (obj_id, slot) = p; heap = sep_heap s in + \obj. sep_map_general obj_id obj {Some slot} s \ object_slots obj = [slot \ cap]" + +definition + sep_any :: "('a \ 'b \ sep_pred) \ ('a \ sep_pred)" where + "sep_any m \ (\p s. \v. (m p v) s)" + +abbreviation "sep_any_map_i \ sep_any sep_map_i" +notation sep_any_map_i ("_ \i -" 76) + +abbreviation "sep_any_map_c \ sep_any sep_map_c" +notation sep_any_map_c ("_ \c -" 76) + +end diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/ex/capDL/Types_D.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Separation_Algebra/ex/capDL/Types_D.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,233 @@ +(* Author: Andrew Boyton, 2012 + Maintainers: Gerwin Klein + Rafal Kolanski +*) + +header "A simplified version of the actual capDL specification." + +theory Types_D +imports "~~/src/HOL/Word/Word" +begin + +(* + * Objects are named by 32 bit words. + * This name may correspond to the memory address of the object. + *) +type_synonym cdl_object_id = "32 word" + +type_synonym cdl_object_set = "cdl_object_id set" + +(* The type we use to represent object sizes. *) +type_synonym cdl_size_bits = nat + +(* An index into a CNode, TCB, or other kernel object that contains caps. *) +type_synonym cdl_cnode_index = nat + +(* A reference to a capability slot. *) +type_synonym cdl_cap_ref = "cdl_object_id \ cdl_cnode_index" + +(* The possible access-control rights that exist in the system. *) +datatype cdl_right = AllowRead | AllowWrite | AllowGrant + + +(* + * Kernel capabilities. + * + * Such capabilities (or "caps") give the holder particular rights to + * a kernel object or system hardware. + * + * Caps have attributes such as the object they point to, the rights + * they give the holder, or how the holder is allowed to interact with + * the target object. + * + * This is a simplified, cut-down version of this datatype for + * demonstration purposes. + *) +datatype cdl_cap = + NullCap + | EndpointCap cdl_object_id "cdl_right set" + | CNodeCap cdl_object_id + | TcbCap cdl_object_id + +(* A mapping from capability identifiers to capabilities. *) +type_synonym cdl_cap_map = "cdl_cnode_index \ cdl_cap option" + +translations + (type) "cdl_cap_map" <= (type) "nat \ cdl_cap option" + (type) "cdl_cap_ref" <= (type) "cdl_object_id \ nat" + +(* A user cap pointer. *) +type_synonym cdl_cptr = "32 word" + +(* Kernel objects *) +record cdl_tcb = + cdl_tcb_caps :: cdl_cap_map + cdl_tcb_fault_endpoint :: cdl_cptr + +record cdl_cnode = + cdl_cnode_caps :: cdl_cap_map + cdl_cnode_size_bits :: cdl_size_bits + +(* + * Kernel objects. + * + * These are in-memory objects that may, over the course of the system + * execution, be created or deleted by users. + * + * Again, a simplified version of the real datatype. + *) +datatype cdl_object = + Endpoint + | Tcb cdl_tcb + | CNode cdl_cnode + +(* + * The current state of the system. + * + * The state record contains the following primary pieces of information: + * + * objects: + * The objects that currently exist in the system. + * + * current_thread: + * The currently running thread. Operations will always be performed + * on behalf of this thread. + * + * ghost_state: (Used for separation logic) + * Which fields are owned by an object. + * In capDL this is all of the fields (or none of them). + * In any concrete state, this will be all of the fields. + *) + + +(* The ghost state tracks which components (fields and slots) are owned by an object. + * Fields + slots are encoded as None + Some nat. + *) +type_synonym cdl_heap = "cdl_object_id \ cdl_object option" +type_synonym cdl_component = "nat option" +type_synonym cdl_components = "cdl_component set" +type_synonym cdl_ghost_state = "cdl_object_id \ cdl_components" + +translations + (type) "cdl_heap" <= (type) "cdl_object_id \ cdl_object option" + (type) "cdl_ghost_state" <= (type) "cdl_object_id \ nat option set" + +record cdl_state = + cdl_objects :: "cdl_heap" + cdl_current_thread :: "cdl_object_id option" + cdl_ghost_state :: "cdl_ghost_state" + + +(* Kernel objects types. *) +datatype cdl_object_type = + EndpointType + | TcbType + | CNodeType + +(* Return the type of an object. *) +definition + object_type :: "cdl_object \ cdl_object_type" +where + "object_type x \ + case x of + Endpoint \ EndpointType + | Tcb _ \ TcbType + | CNode _ \ CNodeType" + +(* + * Getters and setters for various data types. + *) + +(* Capability getters / setters *) + +definition cap_objects :: "cdl_cap \ cdl_object_id set" +where + "cap_objects cap \ + case cap of + TcbCap x \ {x} + | CNodeCap x \ {x} + | EndpointCap x _ \ {x}" + +definition cap_has_object :: "cdl_cap \ bool" +where + "cap_has_object cap \ + case cap of + NullCap \ False + | _ \ True" + +definition cap_object :: "cdl_cap \ cdl_object_id" +where + "cap_object cap \ + if cap_has_object cap + then THE obj_id. cap_objects cap = {obj_id} + else undefined " + +lemma cap_object_simps: + "cap_object (TcbCap x) = x" + "cap_object (CNodeCap x) = x" + "cap_object (EndpointCap x j) = x" + by (simp_all add:cap_object_def cap_objects_def cap_has_object_def) + +definition + cap_rights :: "cdl_cap \ cdl_right set" +where + "cap_rights c \ case c of + EndpointCap _ x \ x + | _ \ UNIV" + +definition + update_cap_rights :: "cdl_right set \ cdl_cap \ cdl_cap" +where + "update_cap_rights r c \ case c of + EndpointCap f1 _ \ EndpointCap f1 r + | _ \ c" + +(* Kernel object getters / setters *) +definition + object_slots :: "cdl_object \ cdl_cap_map" +where + "object_slots obj \ case obj of + CNode x \ cdl_cnode_caps x + | Tcb x \ cdl_tcb_caps x + | _ \ empty" + +definition + update_slots :: "cdl_cap_map \ cdl_object \ cdl_object" +where + "update_slots new_val obj \ case obj of + CNode x \ CNode (x\cdl_cnode_caps := new_val\) + | Tcb x \ Tcb (x\cdl_tcb_caps := new_val\) + | _ \ obj" + +(* Adds new caps to an object. It won't overwrite on a collision. *) +definition + add_to_slots :: "cdl_cap_map \ cdl_object \ cdl_object" +where + "add_to_slots new_val obj \ update_slots (new_val ++ (object_slots obj)) obj" + +definition + slots_of :: "cdl_heap \ cdl_object_id \ cdl_cap_map" +where + "slots_of h \ \obj_id. + case h obj_id of + None \ empty + | Some obj \ object_slots obj" + + +definition + has_slots :: "cdl_object \ bool" +where + "has_slots obj \ case obj of + CNode _ \ True + | Tcb _ \ True + | _ \ False" + +definition + object_at :: "(cdl_object \ bool) \ cdl_object_id \ cdl_heap \ bool" +where + "object_at P p s \ \object. s p = Some object \ P object" + +abbreviation + "ko_at k \ object_at (op = k)" + +end diff -r 77daf1b85cf0 -r a5f5b9336007 Separation_Algebra/sep_tactics.ML --- a/Separation_Algebra/sep_tactics.ML Fri Sep 12 00:47:15 2014 +0800 +++ b/Separation_Algebra/sep_tactics.ML Sat Sep 13 10:07:14 2014 +0800 @@ -121,7 +121,7 @@ (* simp add: sep_conj_ac *) val sep_conj_ac_tac = Simplifier.asm_full_simp_tac - (put_simpset HOL_basic_ss ctxt addsimps SepConj.sep_conj_ac); + (HOL_basic_ss addsimps SepConj.sep_conj_ac); in (* XXX: normally you'd want to keep track of what variables we want to make diff -r 77daf1b85cf0 -r a5f5b9336007 progtut/Advanced.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/progtut/Advanced.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,46 @@ +theory Advanced +imports Essential +begin + +lemma "True" +proof - + ML_prf {* Variable.dest_fixes @{context} *} + fix x y + ML_prf {* Variable.dest_fixes @{context} *} + show ?thesis by auto +qed + +lemma "True" +proof - + fix x y + { fix z w + ML_prf {* Variable.dest_fixes @{context} *} + } + ML_prf {* Variable.dest_fixes @{context} *} + show ?thesis by auto +qed + +ML {* +val ctxt0 = @{context}; +val ([x, y], ctxt1) = Variable.add_fixes ["x", "y"] ctxt0; +val ([z, w], ctxt2) = Variable.add_fixes ["z", "w"] ctxt1; +val trm = @{term "(x, y, z, w)"}; + + pwriteln (Pretty.chunks + [ pretty_term ctxt0 trm, + pretty_term ctxt1 trm, + pretty_term ctxt2 trm ]) +*} + +ML {* + let + val ctxt0 = @{context} + val (y, ctxt1) = Variable.add_fixes ["x"] ctxt0 + val frees = replicate 2 ("x", @{typ nat}) + in + (Variable.variant_frees ctxt0 [] frees, + Variable.variant_frees ctxt1 [] frees, y) + end +*} + +end diff -r 77daf1b85cf0 -r a5f5b9336007 progtut/Advanced.thy~ --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/progtut/Advanced.thy~ Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,2 @@ +theory Advanced +imports Essential \ No newline at end of file diff -r 77daf1b85cf0 -r a5f5b9336007 progtut/Essential.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/progtut/Essential.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,197 @@ +theory Essential +imports FirstStep +begin + +ML {* +fun pretty_helper aux env = + env |> Vartab.dest + |> map aux + |> map (fn (s1, s2) => Pretty.block [s1, Pretty.str " := ", s2]) + |> Pretty.enum "," "[" "]" + |> pwriteln + +fun pretty_tyenv ctxt tyenv = +let + fun get_typs (v, (s, T)) = (TVar (v, s), T) + val print = pairself (pretty_typ ctxt) +in + pretty_helper (print o get_typs) tyenv +end + +fun pretty_env ctxt env = +let + fun get_trms (v, (T, t)) = (Var (v, T), t) + val print = pairself (pretty_term ctxt) +in + pretty_helper (print o get_trms) env +end + +fun prep_trm thy (x, (T, t)) = + (cterm_of thy (Var (x, T)), cterm_of thy t) + +fun prep_ty thy (x, (S, ty)) = + (ctyp_of thy (TVar (x, S)), ctyp_of thy ty) + +fun matcher_inst thy pat trm i = + let + val univ = Unify.matchers thy [(pat, trm)] + val env = nth (Seq.list_of univ) i + val tenv = Vartab.dest (Envir.term_env env) + val tyenv = Vartab.dest (Envir.type_env env) + in + (map (prep_ty thy) tyenv, map (prep_trm thy) tenv) + end +*} + +ML {* + let + val pat = Logic.strip_imp_concl (prop_of @{thm spec}) + val trm = @{term "Trueprop (Q True)"} + val inst = matcher_inst @{theory} pat trm 1 + in + Drule.instantiate_normalize inst @{thm spec} + end +*} + +ML {* +let + val c = Const (@{const_name "plus"}, dummyT) + val o = @{term "1::nat"} + val v = Free ("x", dummyT) +in + Syntax.check_term @{context} (c $ o $ v) +end +*} + +ML {* + val my_thm = + let + val assm1 = @{cprop "\ (x::nat). P x ==> Q x"} + val assm2 = @{cprop "(P::nat => bool) t"} + + val Pt_implies_Qt = + Thm.assume assm1 + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) + |> Thm.forall_elim @{cterm "t::nat"} + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) + + val Qt = Thm.implies_elim Pt_implies_Qt (Thm.assume assm2) + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) + in + Qt + |> Thm.implies_intr assm2 + |> Thm.implies_intr assm1 + end +*} + +local_setup {* + Local_Theory.note ((@{binding "my_thm"}, []), [my_thm]) #> snd +*} + +local_setup {* + Local_Theory.note ((@{binding "my_thm_simp"}, + [Attrib.internal (K Simplifier.simp_add)]), [my_thm]) #> snd +*} + +(* pp 62 *) + +ML {* + Thm.reflexive @{cterm "True"} + |> Simplifier.rewrite_rule [@{thm True_def}] + |> pretty_thm @{context} + |> pwriteln +*} + + +ML {* + val thm = @{thm list.induct} |> forall_intr_vars; + thm |> forall_intr_vars |> cprop_of |> Object_Logic.atomize + |> (fn x => Raw_Simplifier.rewrite_rule [x] thm) +*} + +ML {* +fun atomize_thm thm = +let + val thm' = forall_intr_vars thm + val thm'' = Object_Logic.atomize (cprop_of thm') +in + Raw_Simplifier.rewrite_rule [thm''] thm' +end +*} + +ML {* + @{thm list.induct} |> atomize_thm +*} + +ML {* + Skip_Proof.make_thm @{theory} @{prop "True = False"} +*} + +ML {* +fun pthms_of (PBody {thms, ...}) = map #2 thms +val get_names = map #1 o pthms_of +val get_pbodies = map (Future.join o #3) o pthms_of +*} + +lemma my_conjIa: +shows "A \ B \ A \ B" +apply(rule conjI) +apply(drule conjunct1) +apply(assumption) +apply(drule conjunct2) +apply(assumption) +done + +lemma my_conjIb: +shows "A \ B \ A \ B" +apply(assumption) +done + +lemma my_conjIc: +shows "A \ B \ A \ B" +apply(auto) +done + + +ML {* +@{thm my_conjIa} + |> Thm.proof_body_of + |> get_names +*} + +ML {* +@{thm my_conjIa} + |> Thm.proof_body_of + |> get_pbodies + |> map get_names + |> List.concat +*} + +ML {* +@{thm my_conjIb} + |> Thm.proof_body_of + |> get_pbodies + |> map get_names + |> List.concat +*} + +ML {* +@{thm my_conjIc} + |> Thm.proof_body_of + |> get_pbodies + |> map get_names + |> List.concat +*} + +ML {* +@{thm my_conjIa} + |> Thm.proof_body_of + |> get_pbodies + |> map get_pbodies + |> (map o map) get_names + |> List.concat + |> List.concat + |> duplicates (op=) +*} + +end \ No newline at end of file diff -r 77daf1b85cf0 -r a5f5b9336007 progtut/Essential.thy~ --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/progtut/Essential.thy~ Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,213 @@ +theory Essential +imports FirstStep +begin + +ML {* +fun pretty_helper aux env = + env |> Vartab.dest + |> map aux + |> map (fn (s1, s2) => Pretty.block [s1, Pretty.str " := ", s2]) + |> Pretty.enum "," "[" "]" + |> pwriteln + +fun pretty_tyenv ctxt tyenv = +let + fun get_typs (v, (s, T)) = (TVar (v, s), T) + val print = pairself (pretty_typ ctxt) +in + pretty_helper (print o get_typs) tyenv +end + +fun pretty_env ctxt env = +let + fun get_trms (v, (T, t)) = (Var (v, T), t) + val print = pairself (pretty_term ctxt) +in + pretty_helper (print o get_trms) env +end + +fun prep_trm thy (x, (T, t)) = + (cterm_of thy (Var (x, T)), cterm_of thy t) + +fun prep_ty thy (x, (S, ty)) = + (ctyp_of thy (TVar (x, S)), ctyp_of thy ty) + +fun matcher_inst thy pat trm i = + let + val univ = Unify.matchers thy [(pat, trm)] + val env = nth (Seq.list_of univ) i + val tenv = Vartab.dest (Envir.term_env env) + val tyenv = Vartab.dest (Envir.type_env env) + in + (map (prep_ty thy) tyenv, map (prep_trm thy) tenv) + end +*} + +ML {* + let + val pat = Logic.strip_imp_concl (prop_of @{thm spec}) + val trm = @{term "Trueprop (Q True)"} + val inst = matcher_inst @{theory} pat trm 1 + in + Drule.instantiate_normalize inst @{thm spec} + end +*} + +ML {* +let + val c = Const (@{const_name "plus"}, dummyT) + val o = @{term "1::nat"} + val v = Free ("x", dummyT) +in + Syntax.check_term @{context} (c $ o $ v) +end +*} + +ML {* + val my_thm = + let + val assm1 = @{cprop "\ (x::nat). P x ==> Q x"} + val assm2 = @{cprop "(P::nat => bool) t"} + + val Pt_implies_Qt = + Thm.assume assm1 + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) + |> Thm.forall_elim @{cterm "t::nat"} + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) + + val Qt = Thm.implies_elim Pt_implies_Qt (Thm.assume assm2) + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) + in + Qt + |> Thm.implies_intr assm2 + |> Thm.implies_intr assm1 + end +*} + +local_setup {* + Local_Theory.note ((@{binding "my_thm"}, []), [my_thm]) #> snd +*} + +local_setup {* + Local_Theory.note ((@{binding "my_thm_simp"}, + [Attrib.internal (K Simplifier.simp_add)]), [my_thm]) #> snd +*} + +(* pp 62 *) + +lemma + shows foo_test1: "A \ B \ C" + and foo_test2: "A \ B \ C" sorry + +ML {* + Thm.reflexive @{cterm "True"} + |> Simplifier.rewrite_rule [@{thm True_def}] + |> pretty_thm @{context} + |> pwriteln +*} + +ML {* +Object_Logic.rulify @{thm foo_test2} +*} + + +ML {* + val thm = @{thm foo_test1}; + thm + |> cprop_of + |> Object_Logic.atomize + |> (fn x => Raw_Simplifier.rewrite_rule [x] thm) +*} + +ML {* + val thm = @{thm list.induct} |> forall_intr_vars; + thm |> forall_intr_vars |> cprop_of |> Object_Logic.atomize + |> (fn x => Raw_Simplifier.rewrite_rule [x] thm) +*} + +ML {* +fun atomize_thm thm = +let + val thm' = forall_intr_vars thm + val thm'' = Object_Logic.atomize (cprop_of thm') +in + Raw_Simplifier.rewrite_rule [thm''] thm' +end +*} + +ML {* + @{thm list.induct} |> atomize_thm +*} + +ML {* + Skip_Proof.make_thm @{theory} @{prop "True = False"} +*} + +ML {* +fun pthms_of (PBody {thms, ...}) = map #2 thms +val get_names = map #1 o pthms_of +val get_pbodies = map (Future.join o #3) o pthms_of +*} + +lemma my_conjIa: +shows "A \ B \ A \ B" +apply(rule conjI) +apply(drule conjunct1) +apply(assumption) +apply(drule conjunct2) +apply(assumption) +done + +lemma my_conjIb: +shows "A \ B \ A \ B" +apply(assumption) +done + +lemma my_conjIc: +shows "A \ B \ A \ B" +apply(auto) +done + + +ML {* +@{thm my_conjIa} + |> Thm.proof_body_of + |> get_names +*} + +ML {* +@{thm my_conjIa} + |> Thm.proof_body_of + |> get_pbodies + |> map get_names + |> List.concat +*} + +ML {* +@{thm my_conjIb} + |> Thm.proof_body_of + |> get_pbodies + |> map get_names + |> List.concat +*} + +ML {* +@{thm my_conjIc} + |> Thm.proof_body_of + |> get_pbodies + |> map get_names + |> List.concat +*} + +ML {* +@{thm my_conjIa} + |> Thm.proof_body_of + |> get_pbodies + |> map get_pbodies + |> (map o map) get_names + |> List.concat + |> List.concat + |> duplicates (op=) +*} + +end \ No newline at end of file diff -r 77daf1b85cf0 -r a5f5b9336007 progtut/FirstStep.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/progtut/FirstStep.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,173 @@ +theory FirstStep +imports Main +begin + +ML {* +val pretty_term = Syntax.pretty_term +val pwriteln = Pretty.writeln +fun pretty_terms ctxt trms = + Pretty.block (Pretty.commas (map (pretty_term ctxt) trms)) +val show_type_ctxt = Config.put show_types true @{context} +val show_all_types_ctxt = Config.put show_all_types true @{context} +fun pretty_cterm ctxt ctrm = + pretty_term ctxt (term_of ctrm) +fun pretty_thm ctxt thm = + pretty_term ctxt (prop_of thm) +fun pretty_thm_no_vars ctxt thm = +let + val ctxt' = Config.put show_question_marks false ctxt +in + pretty_term ctxt' (prop_of thm) +end +fun pretty_thms ctxt thms = + Pretty.block (Pretty.commas (map (pretty_thm ctxt) thms)) +fun pretty_thms_no_vars ctxt thms = + Pretty.block (Pretty.commas (map (pretty_thm_no_vars ctxt) thms)) +fun pretty_typ ctxt ty = Syntax.pretty_typ ctxt ty +fun pretty_typs ctxt tys = + Pretty.block (Pretty.commas (map (pretty_typ ctxt) tys)) +fun pretty_ctyp ctxt cty = pretty_typ ctxt (typ_of cty) +fun pretty_ctyps ctxt ctys = + Pretty.block (Pretty.commas (map (pretty_ctyp ctxt) ctys)) +fun `f = fn x => (f x, x) +fun (x, y) |>> f = (f x, y) +fun (x, y) ||> f = (x, f y) +fun (x, y) |-> f = f x y +*} + +ML {* + val _ = pretty_term @{context} @{term "x + y + z"} |> pwriteln + val _ = pretty_terms @{context} [@{term "x + y"}, @{term "x + y + z"}] |> pwriteln + val show_type_ctxt = Config.put show_types true @{context} + *} + +ML {* + pwriteln (pretty_term show_type_ctxt @{term "(1::nat, x)"}) +*} + +ML {* +pwriteln (Pretty.str (cat_lines ["foo", "bar"])) +*} + +ML {* + fun apply_fresh_args f ctxt = + f |> fastype_of + |> binder_types + |> map (pair "z") + |> Variable.variant_frees ctxt [f] + |> map Free + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) + |> curry list_comb f +*} + +ML {* +let + val trm = @{term "P::nat => int => unit => bool"} + val ctxt = ML_Context.the_local_context () +in + apply_fresh_args trm ctxt + |> pretty_term ctxt + |> pwriteln +end +*} + +ML {* + fun inc_by_three x = + x |> (fn x => x + 1) + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) + |> (fn x => x + 2) +*} + +ML {* + fun `f = fn x => (f x, x) +*} + +ML {* + fun inc_as_pair x = + x |> `(fn x => x + 1) + |> (fn (x, y) => (x, y + 1)) +*} + +ML {* + 3 |> inc_as_pair +*} + +ML {* + fun acc_incs x = + x |> (fn x => (0, x)) + ||>> (fn x => (x, x + 1)) + ||>> (fn x => (x, x + 1)) + ||>> (fn x => (x, x + 1)) +*} + +ML {* + 2 |> acc_incs +*} + +ML {* +let + val ((names1, names2), _) = + @{context} + |> Variable.variant_fixes (replicate 4 "x") + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) + ||>> Variable.variant_fixes (replicate 5 "x") + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) +in + (names1, names2) +end +*} + +ML {* + @{context} + |> Local_Defs.add_def ((@{binding "One"}, NoSyn), @{term "1::nat"}) + ||>> Local_Defs.add_def ((@{binding "Two"}, NoSyn), @{term "2::nat"}) + ||>> Local_Defs.add_def ((@{binding "Three"}, NoSyn), @{term "3::nat"}) +*} + +ML {* +let + val ctxt = @{context} +in + Syntax.parse_term ctxt "m + n" + |> singleton (Syntax.check_terms ctxt) + |> pretty_term ctxt + |> pwriteln +end +*} + +ML {* + val term_pat_setup = + let + val parser = Args.context -- Scan.lift Args.name_source + fun term_pat (ctxt, str) = + str |> Proof_Context.read_term_pattern ctxt + |> ML_Syntax.print_term + |> ML_Syntax.atomic + in + ML_Antiquote.inline @{binding "term_pat"} (parser >> term_pat) + end +*} + +setup {* term_pat_setup *} + + +ML {* +val type_pat_setup = +let + val parser = Args.context -- Scan.lift Args.name_source + fun typ_pat (ctxt, str) = + let + val ctxt' = Proof_Context.set_mode Proof_Context.mode_schematic ctxt + in + str |> Syntax.read_typ ctxt' + |> ML_Syntax.print_typ + |> ML_Syntax.atomic + end +in + ML_Antiquote.inline @{binding "typ_pat"} (parser >> typ_pat) +end +*} + +setup {* type_pat_setup *} + +end \ No newline at end of file diff -r 77daf1b85cf0 -r a5f5b9336007 progtut/FirstStep.thy~ --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/progtut/FirstStep.thy~ Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,173 @@ +theory FistStep +imports Main +begin + +ML {* +val pretty_term = Syntax.pretty_term +val pwriteln = Pretty.writeln +fun pretty_terms ctxt trms = + Pretty.block (Pretty.commas (map (pretty_term ctxt) trms)) +val show_type_ctxt = Config.put show_types true @{context} +val show_all_types_ctxt = Config.put show_all_types true @{context} +fun pretty_cterm ctxt ctrm = + pretty_term ctxt (term_of ctrm) +fun pretty_thm ctxt thm = + pretty_term ctxt (prop_of thm) +fun pretty_thm_no_vars ctxt thm = +let + val ctxt' = Config.put show_question_marks false ctxt +in + pretty_term ctxt' (prop_of thm) +end +fun pretty_thms ctxt thms = + Pretty.block (Pretty.commas (map (pretty_thm ctxt) thms)) +fun pretty_thms_no_vars ctxt thms = + Pretty.block (Pretty.commas (map (pretty_thm_no_vars ctxt) thms)) +fun pretty_typ ctxt ty = Syntax.pretty_typ ctxt ty +fun pretty_typs ctxt tys = + Pretty.block (Pretty.commas (map (pretty_typ ctxt) tys)) +fun pretty_ctyp ctxt cty = pretty_typ ctxt (typ_of cty) +fun pretty_ctyps ctxt ctys = + Pretty.block (Pretty.commas (map (pretty_ctyp ctxt) ctys)) +fun `f = fn x => (f x, x) +fun (x, y) |>> f = (f x, y) +fun (x, y) ||> f = (x, f y) +fun (x, y) |-> f = f x y +*} + +ML {* + val _ = pretty_term @{context} @{term "x + y + z"} |> pwriteln + val _ = pretty_terms @{context} [@{term "x + y"}, @{term "x + y + z"}] |> pwriteln + val show_type_ctxt = Config.put show_types true @{context} + *} + +ML {* + pwriteln (pretty_term show_type_ctxt @{term "(1::nat, x)"}) +*} + +ML {* +pwriteln (Pretty.str (cat_lines ["foo", "bar"])) +*} + +ML {* + fun apply_fresh_args f ctxt = + f |> fastype_of + |> binder_types + |> map (pair "z") + |> Variable.variant_frees ctxt [f] + |> map Free + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) + |> curry list_comb f +*} + +ML {* +let + val trm = @{term "P::nat => int => unit => bool"} + val ctxt = ML_Context.the_local_context () +in + apply_fresh_args trm ctxt + |> pretty_term ctxt + |> pwriteln +end +*} + +ML {* + fun inc_by_three x = + x |> (fn x => x + 1) + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) + |> (fn x => x + 2) +*} + +ML {* + fun `f = fn x => (f x, x) +*} + +ML {* + fun inc_as_pair x = + x |> `(fn x => x + 1) + |> (fn (x, y) => (x, y + 1)) +*} + +ML {* + 3 |> inc_as_pair +*} + +ML {* + fun acc_incs x = + x |> (fn x => (0, x)) + ||>> (fn x => (x, x + 1)) + ||>> (fn x => (x, x + 1)) + ||>> (fn x => (x, x + 1)) +*} + +ML {* + 2 |> acc_incs +*} + +ML {* +let + val ((names1, names2), _) = + @{context} + |> Variable.variant_fixes (replicate 4 "x") + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) + ||>> Variable.variant_fixes (replicate 5 "x") + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) +in + (names1, names2) +end +*} + +ML {* + @{context} + |> Local_Defs.add_def ((@{binding "One"}, NoSyn), @{term "1::nat"}) + ||>> Local_Defs.add_def ((@{binding "Two"}, NoSyn), @{term "2::nat"}) + ||>> Local_Defs.add_def ((@{binding "Three"}, NoSyn), @{term "3::nat"}) +*} + +ML {* +let + val ctxt = @{context} +in + Syntax.parse_term ctxt "m + n" + |> singleton (Syntax.check_terms ctxt) + |> pretty_term ctxt + |> pwriteln +end +*} + +ML {* + val term_pat_setup = + let + val parser = Args.context -- Scan.lift Args.name_source + fun term_pat (ctxt, str) = + str |> Proof_Context.read_term_pattern ctxt + |> ML_Syntax.print_term + |> ML_Syntax.atomic + in + ML_Antiquote.inline @{binding "term_pat"} (parser >> term_pat) + end +*} + +setup {* term_pat_setup *} + + +ML {* +val type_pat_setup = +let + val parser = Args.context -- Scan.lift Args.name_source + fun typ_pat (ctxt, str) = + let + val ctxt' = Proof_Context.set_mode Proof_Context.mode_schematic ctxt + in + str |> Syntax.read_typ ctxt' + |> ML_Syntax.print_typ + |> ML_Syntax.atomic + end +in + ML_Antiquote.inline @{binding "typ_pat"} (parser >> typ_pat) +end +*} + +setup {* type_pat_setup *} + +end \ No newline at end of file diff -r 77daf1b85cf0 -r a5f5b9336007 progtut/Tactical.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/progtut/Tactical.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,64 @@ +theory Tactical +imports Advanced +begin + +ML {* +fun my_print_tac ctxt thm = +let + val _ = tracing (Pretty.string_of (pretty_thm_no_vars ctxt thm)) +in + Seq.single thm +end +*} + +ML {* +let + val ctxt = @{context} + val goal = @{prop "P \ Q \ Q \ P"} +in + Goal.prove ctxt ["x", "y", "u", "v", "P", "Q"] [@{prop "x = y"}, @{prop "u = v"}] goal + (fn {prems = pms, context = ctxt} => + (my_print_tac ctxt) + THEN etac @{thm disjE} 1 THEN + (my_print_tac ctxt) + THEN rtac @{thm disjI2} 1 + THEN (my_print_tac ctxt) + THEN atac 1 + THEN rtac @{thm disjI1} 1 + THEN atac 1) +end +*} + +ML {* + Goal.prove +*} + +ML {* +fun pretty_cterms ctxt ctrms = + Pretty.block (Pretty.commas (map (pretty_cterm ctxt) ctrms)) +*} + +ML {* +fun foc_tac {prems, params, asms, concl, context, schematics} = +let + fun assgn1 f1 f2 xs = + let + val out = map (fn (x, y) => Pretty.enum ":=" "" "" [f1 x, f2 y]) xs + in + Pretty.list "" "" out + end + fun assgn2 f xs = assgn1 f f xs + val pps = map (fn (s, pp) => Pretty.block [Pretty.str s, pp]) + [("params: ", assgn1 Pretty.str (pretty_cterm context) params), + ("assumptions: ", pretty_cterms context asms), + ("conclusion: ", pretty_cterm context concl), + ("premises: ", pretty_thms_no_vars context prems), + ("schematics: ", assgn2 (pretty_cterm context) (snd schematics))] +in + tracing (Pretty.string_of (Pretty.chunks pps)); all_tac +end +*} + +notation (output) "prop" ("#_" [1000] 1000) + +end diff -r 77daf1b85cf0 -r a5f5b9336007 progtut/Tactical.thy~ --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/progtut/Tactical.thy~ Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,76 @@ +theory Tactical +imports Advanced +begin + + + + +ML {* +fun my_print_tac ctxt thm = +let + val _ = tracing (Pretty.string_of (pretty_thm_no_vars ctxt thm)) +in + Seq.single thm +end +*} + +ML {* +let + val ctxt = @{context} + val goal = @{prop "P \ Q \ Q \ P"} +in + Goal.prove ctxt ["x", "y", "u", "v", "P", "Q"] [@{prop "x = y"}, @{prop "u = v"}] goal + (fn {prems = pms, context = ctxt} => + (my_print_tac ctxt) + THEN etac @{thm disjE} 1 THEN + (my_print_tac ctxt) + THEN rtac @{thm disjI2} 1 + THEN (my_print_tac ctxt) + THEN atac 1 + THEN rtac @{thm disjI1} 1 + THEN atac 1) +end +*} + +ML {* + Goal.prove +*} + +ML {* +fun pretty_cterms ctxt ctrms = + Pretty.block (Pretty.commas (map (pretty_cterm ctxt) ctrms)) +*} + +ML {* +fun foc_tac {prems, params, asms, concl, context, schematics} = +let + fun assgn1 f1 f2 xs = + let + val out = map (fn (x, y) => Pretty.enum ":=" "" "" [f1 x, f2 y]) xs + in + Pretty.list "" "" out + end + fun assgn2 f xs = assgn1 f f xs + val pps = map (fn (s, pp) => Pretty.block [Pretty.str s, pp]) + [("params: ", assgn1 Pretty.str (pretty_cterm context) params), + ("assumptions: ", pretty_cterms context asms), + ("conclusion: ", pretty_cterm context concl), + ("premises: ", pretty_thms_no_vars context prems), + ("schematics: ", assgn2 (pretty_cterm context) (snd schematics))] +in + tracing (Pretty.string_of (Pretty.chunks pps)); all_tac +end +*} + + +notation (output) "prop" ("#_" [1000] 1000) + +lemma + shows " \A; B\ \ A \ B" +apply(tactic {* my_print_tac @{context} *}) +apply(rule conjI) +apply(tactic {* my_print_tac @{context} *}) +apply(assumption) +apply(tactic {* my_print_tac @{context} *}) +apply(assumption) +apply(tactic {* my_print_tac @{context} *}) diff -r 77daf1b85cf0 -r a5f5b9336007 progtut/progtut.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/progtut/progtut.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,173 @@ +theory FistStep +imports Main +begin + +ML {* +val pretty_term = Syntax.pretty_term +val pwriteln = Pretty.writeln +fun pretty_terms ctxt trms = + Pretty.block (Pretty.commas (map (pretty_term ctxt) trms)) +val show_type_ctxt = Config.put show_types true @{context} +val show_all_types_ctxt = Config.put show_all_types true @{context} +fun pretty_cterm ctxt ctrm = + pretty_term ctxt (term_of ctrm) +fun pretty_thm ctxt thm = + pretty_term ctxt (prop_of thm) +fun pretty_thm_no_vars ctxt thm = +let + val ctxt' = Config.put show_question_marks false ctxt +in + pretty_term ctxt' (prop_of thm) +end +fun pretty_thms ctxt thms = + Pretty.block (Pretty.commas (map (pretty_thm ctxt) thms)) +fun pretty_thms_no_vars ctxt thms = + Pretty.block (Pretty.commas (map (pretty_thm_no_vars ctxt) thms)) +fun pretty_typ ctxt ty = Syntax.pretty_typ ctxt ty +fun pretty_typs ctxt tys = + Pretty.block (Pretty.commas (map (pretty_typ ctxt) tys)) +fun pretty_ctyp ctxt cty = pretty_typ ctxt (typ_of cty) +fun pretty_ctyps ctxt ctys = + Pretty.block (Pretty.commas (map (pretty_ctyp ctxt) ctys)) +fun `f = fn x => (f x, x) +fun (x, y) |>> f = (f x, y) +fun (x, y) ||> f = (x, f y) +fun (x, y) |-> f = f x y +*} + +ML {* + val _ = pretty_term @{context} @{term "x + y + z"} |> pwriteln + val _ = pretty_terms @{context} [@{term "x + y"}, @{term "x + y + z"}] |> pwriteln + val show_type_ctxt = Config.put show_types true @{context} + *} + +ML {* + pwriteln (pretty_term show_type_ctxt @{term "(1::nat, x)"}) +*} + +ML {* +pwriteln (Pretty.str (cat_lines ["foo", "bar"])) +*} + +ML {* + fun apply_fresh_args f ctxt = + f |> fastype_of + |> binder_types + |> map (pair "z") + |> Variable.variant_frees ctxt [f] + |> map Free + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) + |> curry list_comb f +*} + +ML {* +let + val trm = @{term "P::nat => int => unit => bool"} + val ctxt = ML_Context.the_local_context () +in + apply_fresh_args trm ctxt + |> pretty_term ctxt + |> pwriteln +end +*} + +ML {* + fun inc_by_three x = + x |> (fn x => x + 1) + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) + |> (fn x => x + 2) +*} + +ML {* + fun `f = fn x => (f x, x) +*} + +ML {* + fun inc_as_pair x = + x |> `(fn x => x + 1) + |> (fn (x, y) => (x, y + 1)) +*} + +ML {* + 3 |> inc_as_pair +*} + +ML {* + fun acc_incs x = + x |> (fn x => (0, x)) + ||>> (fn x => (x, x + 1)) + ||>> (fn x => (x, x + 1)) + ||>> (fn x => (x, x + 1)) +*} + +ML {* + 2 |> acc_incs +*} + +ML {* +let + val ((names1, names2), _) = + @{context} + |> Variable.variant_fixes (replicate 4 "x") + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) + ||>> Variable.variant_fixes (replicate 5 "x") + |> tap (fn x => pwriteln (Pretty.str (@{make_string} x))) +in + (names1, names2) +end +*} + +ML {* + @{context} + |> Local_Defs.add_def ((@{binding "One"}, NoSyn), @{term "1::nat"}) + ||>> Local_Defs.add_def ((@{binding "Two"}, NoSyn), @{term "2::nat"}) + ||>> Local_Defs.add_def ((@{binding "Three"}, NoSyn), @{term "3::nat"}) +*} + +ML {* +let + val ctxt = @{context} +in + Syntax.parse_term ctxt "m + n" + |> singleton (Syntax.check_terms ctxt) + |> pretty_term ctxt + |> pwriteln +end +*} + +ML {* + val term_pat_setup = + let + val parser = Args.context -- Scan.lift Args.name_source + fun term_pat (ctxt, str) = + str |> Proof_Context.read_term_pattern ctxt + |> ML_Syntax.print_term + |> ML_Syntax.atomic + in + ML_Antiquote.inline @{binding "term_pat"} (parser >> term_pat) + end +*} + +setup {* term_pat_setup *} + + +ML {* +val type_pat_setup = +let + val parser = Args.context -- Scan.lift Args.name_source + fun typ_pat (ctxt, str) = + let + val ctxt' = Proof_Context.set_mode Proof_Context.mode_schematic ctxt + in + str |> Syntax.read_typ ctxt' + |> ML_Syntax.print_typ + |> ML_Syntax.atomic + end +in + ML_Antiquote.inline @{binding "typ_pat"} (parser >> typ_pat) +end +*} + +setup {* type_pat_setup *} + +end \ No newline at end of file diff -r 77daf1b85cf0 -r a5f5b9336007 progtut/progtut.thy~ diff -r 77daf1b85cf0 -r a5f5b9336007 progtut/progtutorial.pdf Binary file progtut/progtutorial.pdf has changed diff -r 77daf1b85cf0 -r a5f5b9336007 thys/ROOT --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys/ROOT Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,24 @@ +session "Hoare_gen" = "HOL" + + options [document = pdf] + theories [document = false] + Hoare_gen + +session "Hoare_tm_basis" = "Hoare_gen" + + options [document = pdf] + theories [document = false, document_output = "output_tm", document_variants = "hoare_tm"] + Hoare_tm_basis + files + "document/root_tm.tex" + +session "Hoare_tm" = "Hoare_tm_basis" + + options [document = pdf] + theories [document = false, document_output = "output_tm", document_variants = "hoare_tm"] + Hoare_tm + files + "document/root_tm.tex" + +session "Hoare_abc" = "Hoare_tm" + + options [document = pdf] + theories [document = false, document_output = "./output_abc"] + Hoare_abc + diff -r 77daf1b85cf0 -r a5f5b9336007 thys2/Data_slot.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/Data_slot.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,25 @@ +theory Data_slot +imports Main +begin + +ML {* + fun ML_exec prog = ML_Compiler.eval true Position.none + (ML_Lex.tokenize prog) + + fun prog_data_slot name tp empty = + "structure " ^ name ^ " = Theory_Data ( " ^ + "type T = "^ tp ^" " ^ + "val empty = " ^ empty ^ " ; "^ + "val extend = I; "^ + "fun merge (x, y) = x) "^ + "fun "^ name ^ "_store v = Context.>>(Context.map_theory (" ^ name ^".map (fn x => v))) "^ + "fun "^ name ^ "_get () = Context.>>>(Context.map_theory_result (fn th => "^ + " ("^ name ^".get th, th))) " + + fun data_slot name tp empty = ML_exec (prog_data_slot name tp empty) + val empty_term = @{term "x"} + val empty_cterm = @{cterm "x"} + val empty_thm = @{thm ext} +*} + +end diff -r 77daf1b85cf0 -r a5f5b9336007 thys2/Hoare_abc.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/Hoare_abc.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,1223 @@ +header {* + {\em Abacus} defined as macros of TM + *} + +theory Hoare_abc +imports Hoare_tm Finite_Set +begin + + +text {* + {\em Abacus} instructions +*} + +(* +text {* The following Abacus instructions will be replaced by TM macros. *} +datatype abc_inst = + -- {* @{text "Inc n"} increments the memory cell (or register) + with address @{text "n"} by one. + *} + Inc nat + -- {* + @{text "Dec n label"} decrements the memory cell with address @{text "n"} by one. + If cell @{text "n"} is already zero, no decrements happens and the executio jumps to + the instruction labeled by @{text "label"}. + *} + | Dec nat nat + -- {* + @{text "Goto label"} unconditionally jumps to the instruction labeled by @{text "label"}. + *} + | Goto nat + +*) + +datatype aresource = + M nat nat + (* | C nat abc_inst *) (* C resource is not needed because there is no Abacus code any more *) + | At nat + | Faults nat + + +section {* An interpretation from Abacus to Turing Machine *} + +fun recse_map :: "nat list \ aresource \ tassert" where + "recse_map ks (M a v) = <(a < length ks \ ks!a = v \ a \ length ks \ v = 0)>" | + "recse_map ks (At l) = st l" | + "recse_map ks (Faults n) = sg {TFaults n}" + +definition "IA ars = (EXS ks i. ps 2 \* zero 0 \* zero 1 \* (reps 2 i ks) \* + fam_conj {i<..} zero \* + fam_conj ars (recse_map ks))" + + +section {* A virtually defined Abacus *} + +text {* The following Abacus instructions are to be defined as TM macros *} + +definition "pc l = sg {At l}" + +definition "mm a v =sg ({M a v})" + +type_synonym assert = "aresource set \ bool" + +lemma tm_hoare_inc1: + assumes h: "a < length ks \ ks ! a = v \ length ks \ a \ v = 0" + shows " + \st i \* ps u \* zero (u - 2) \* zero (u - 1) \* reps u ia ks \* fam_conj {ia<..} zero\ + i :[Inc a ]: j + \st j \* + ps u \* + zero (u - 2) \* + zero (u - 1) \* + reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1) ((list_ext a ks)[a := Suc v]) \* + fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1<..} zero\" + using h +proof + assume hh: "a < length ks \ ks ! a = v" + hence "a < length ks" by simp + from list_ext_lt [OF this] tm_hoare_inc00[OF hh] + show ?thesis by simp +next + assume "length ks \ a \ v = 0" + from tm_hoare_inc01[OF this] + show ?thesis by simp +qed + +lemma tm_hoare_inc2: + assumes "mm a v sr" + shows " + \ (fam_conj sr (recse_map ks) \* + st i \* ps 2 \* zero 0 \* zero 1 \* reps 2 ia ks \* fam_conj {ia<..} zero) \ + i:[ (Inc a) ]:j + \ (fam_conj {M a (Suc v)} (recse_map (list_ext a ks[a := Suc v])) \* + st j \* + ps 2 \* + zero 0 \* + zero 1 \* + reps 2 (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1) ((list_ext a ks)[a := Suc v]) \* + fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1<..} zero)\" +proof - + from `mm a v sr` have eq_sr: "sr = {M a v}" by (auto simp:mm_def sg_def) + from tm_hoare_inc1[where u = 2] + have "a < length ks \ ks ! a = v \ length ks \ a \ v = 0 \ + \st i \* ps 2 \* zero 0 \* zero 1 \* reps 2 ia ks \* fam_conj {ia<..} zero\ + i :[Inc a ]: j + \(st j \* + ps 2 \* + zero 0 \* + zero 1 \* + reps 2 (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1) ((list_ext a ks)[a := Suc v]) \* + fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1<..} zero)\" by simp + thus ?thesis + apply (unfold eq_sr) + apply (simp add:fam_conj_simps list_ext_get_upd list_ext_len) + by (rule tm.pre_condI, blast) +qed + +locale IA_disjoint = + fixes s s' s1 cnf + assumes h_IA: "IA (s + s') s1" + and h_disj: "s ## s'" + and h_conf: "s1 \ trset_of cnf" +begin + +lemma at_disj1: + assumes at_in: "At i \ s" + shows "At j \ s'" +proof + from h_IA[unfolded IA_def] + obtain ks idx + where "((ps 2 \* zero 0 \* zero 1 \* + reps 2 idx ks \* fam_conj {idx<..} zero) \* + fam_conj (s + s') (recse_map ks)) s1" (is "(?P \* ?Q) s1") + by (auto elim!:EXS_elim simp:sep_conj_ac) + then obtain ss1 ss2 where "s1 = ss1 + ss2" "ss1 ## ss2" "?P ss1" "?Q ss2" + by (auto elim:sep_conjE) + from `?Q ss2`[unfolded fam_conj_disj_simp[OF h_disj]] + obtain tt1 tt2 + where "ss2 = tt1 + tt2" "tt1 ## tt2" + "(fam_conj s (recse_map ks)) tt1" + "(fam_conj s' (recse_map ks)) tt2" + by (auto elim:sep_conjE) + assume "At j \ s'" + from `(fam_conj s' (recse_map ks)) tt2` [unfolded fam_conj_elm_simp[OF this]] + `ss2 = tt1 + tt2` `s1 = ss1 + ss2` h_conf + have "TAt j \ trset_of cnf" + by (auto elim!:sep_conjE simp: st_def sg_def tpn_set_def set_ins_def) + moreover from `(fam_conj s (recse_map ks)) tt1` [unfolded fam_conj_elm_simp[OF at_in]] + `ss2 = tt1 + tt2` `s1 = ss1 + ss2` h_conf + have "TAt i \ trset_of cnf" + by (auto elim!:sep_conjE simp: st_def sg_def tpn_set_def set_ins_def) + ultimately have "i = j" + by (cases cnf, simp add:trset_of.simps tpn_set_def) + from at_in `At j \ s'` h_disj + show False + by (unfold `i = j`, auto simp:set_ins_def) +qed + +lemma at_disj2: "At i \ s' \ At j \ s" + by (metis at_disj1) + +lemma m_disj1: + assumes m_in: "M a v \ s" + shows "M a v' \ s'" +proof + from h_IA[unfolded IA_def] + obtain ks idx + where "((ps 2 \* zero 0 \* zero 1 \* + reps 2 idx ks \* fam_conj {idx<..} zero) \* + fam_conj (s + s') (recse_map ks)) s1" (is "(?P \* ?Q) s1") + by (auto elim!:EXS_elim simp:sep_conj_ac) + then obtain ss1 ss2 where "s1 = ss1 + ss2" "ss1 ## ss2" "?P ss1" "?Q ss2" + by (auto elim:sep_conjE) + from `?Q ss2`[unfolded fam_conj_disj_simp[OF h_disj]] + obtain tt1 tt2 + where "ss2 = tt1 + tt2" "tt1 ## tt2" + "(fam_conj s (recse_map ks)) tt1" + "(fam_conj s' (recse_map ks)) tt2" + by (auto elim:sep_conjE) + assume "M a v' \ s'" + from `(fam_conj s' (recse_map ks)) tt2` [unfolded fam_conj_elm_simp[OF this] + recse_map.simps] + have "(a < length ks \ ks ! a = v' \ length ks \ a \ v' = 0)" + by (auto simp:pasrt_def) + moreover from `(fam_conj s (recse_map ks)) tt1` [unfolded fam_conj_elm_simp[OF m_in] + recse_map.simps] + have "a < length ks \ ks ! a = v \ length ks \ a \ v = 0" + by (auto simp:pasrt_def) + moreover note m_in `M a v' \ s'` h_disj + ultimately show False + by (auto simp:set_ins_def) +qed + +lemma m_disj2: "M a v \ s' \ M a v' \ s" + by (metis m_disj1) + +end + +lemma EXS_elim1: + assumes "((EXS x. P(x)) \* r) s" + obtains x where "(P(x) \* r) s" + by (metis EXS_elim assms sep_conj_exists1) + +lemma hoare_inc[step]: "IA. \ pc i ** mm a v \ + i:[ (Inc a) ]:j + \ pc j ** mm a (Suc v)\" + (is "IA. \ pc i ** ?P \ + i:[ ?code ?e ]:j + \ pc ?e ** ?Q\") +proof(induct rule:tm.IHoareI) + case (IPre s' s r cnf) + let ?cnf = "(trset_of cnf)" + from IPre + have h: "(pc i \* ?P) s" "(s ## s')" "(IA (s + s') \* i :[ ?code(?e) ]: j \* r) ?cnf" + by (metis condD)+ + from h(1) obtain sr where + eq_s: "s = {At i} \ sr" "{At i} ## sr" "?P sr" + by (auto dest!:sep_conjD simp:set_ins_def pc_def sg_def) + hence "At i \ s" by auto + from h(3) obtain s1 s2 s3 + where hh: "?cnf = s1 + s2 + s3" + "s1 ## s2 \ s2 ## s3 \ s1 ## s3" + "IA (s + s') s1" + "(i :[ ?code ?e ]: j) s2" + "r s3" + apply (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac) + by (metis sep_add_commute sep_disj_addD1 sep_disj_addD2 sep_disj_commuteI) + interpret ia_disj: IA_disjoint s s' s1 cnf + proof + from `IA (s + s') s1` show "IA (s + s') s1" . + next + from `s ## s'` show "s ## s'" . + next + from hh(1) show "s1 \ ?cnf" by (auto simp:set_ins_def) + qed + from hh(1) have s1_belong: "s1 \ ?cnf" by (auto simp:set_ins_def) + from hh(3) + have "(EXS ks ia. + ps 2 \* + zero 0 \* + zero 1 \* + reps 2 ia ks \* + fam_conj {ia<..} zero \* + (st i \* fam_conj sr (recse_map ks)) \* fam_conj s' (recse_map ks)) + s1" + apply (unfold IA_def fam_conj_disj_simp[OF h(2)]) + apply (unfold eq_s) + by (insert eq_s(2), simp add: fam_conj_disj_simp[OF eq_s(2)] fam_conj_simps set_ins_def) + then obtain ks ia + where "((fam_conj sr (recse_map ks) \* st i \* ps 2 \* zero 0 \* zero 1 \* + reps 2 ia ks \* fam_conj {ia<..} zero) \* fam_conj s' (recse_map ks)) s1" + (is "(?PP \* ?QQ) s1") + by (unfold pred_ex_def, auto simp:sep_conj_ac) + then obtain ss1 ss2 where pres: + "s1 = ss1 + ss2" "ss1 ## ss2" + "?PP ss1" + "?QQ ss2" + by (auto elim!:sep_conjE intro!:sep_conjI) + from ia_disj.at_disj1 [OF `At i \ s`] + have at_fresh_s': "At ?e \ s'" . + have at_fresh_sr: "At ?e \ sr" + proof + assume at_in: "At ?e \ sr" + from h(3)[unfolded IA_def fam_conj_disj_simp[OF h(2)]] + have "TAt ?e \ trset_of cnf" + apply (elim EXS_elim1) + apply (unfold eq_s fam_conj_disj_simp[OF eq_s(2), unfolded set_ins_def] + fam_conj_elm_simp[OF at_in]) + apply (erule_tac sep_conjE, unfold set_ins_def)+ + by (auto simp:sep_conj_ac fam_conj_simps st_def sg_def tpn_set_def) + moreover from pres(1, 3) hh(1) have "TAt i \ trset_of cnf" + apply(erule_tac sep_conjE) + apply(erule_tac sep_conjE) + by (auto simp:st_def tpc_set_def sg_def set_ins_def) + ultimately have "i = ?e" + by (cases cnf, auto simp:tpn_set_def trset_of.simps) + from eq_s[unfolded this] at_in + show "False" by (auto simp:set_ins_def) + qed + from pres(3) and hh(2, 4, 5) pres(2, 4) + have pres1: "(?PP \* i :[ ?code(?e)]: j \* (r \* (fam_conj s' (recse_map ks)))) + (trset_of cnf)" + apply (unfold hh(1) pres(1)) + apply (rule sep_conjI[where x=ss1 and y = "ss2 + s2 + s3"], assumption+) + apply (rule sep_conjI[where x="s2" and y = "ss2 + s3"], assumption+) + apply (rule sep_conjI[where x="s3" and y = ss2], assumption+) + by (auto simp:set_ins_def) + (*****************************************************************************) + let ?ks_f = "\ sr ks. list_ext a ks[a := Suc v]" + let ?elm_f = "\ sr. {M a (Suc v)}" + let ?idx_f = "\ sr ks ia. (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1)" + (*----------------------------------------------------------------------------*) + (******************************************************************************) + from tm_hoare_inc2 [OF eq_s(3), unfolded tm.Hoare_gen_def, rule_format, OF pres1] + obtain k where "((fam_conj (?elm_f sr) (recse_map (?ks_f sr ks)) \* + st ?e \* + ps 2 \* + zero 0 \* zero 1 \* reps 2 (?idx_f sr ks ia) (?ks_f sr ks) \* + fam_conj {?idx_f sr ks ia<..} zero) \* + i :[ ?code ?e ]: j \* r \* fam_conj s' (recse_map ks)) + (trset_of (sep_exec.run tstep (Suc k) cnf))" by blast + (*----------------------------------------------------------------------------*) + moreover have "(pc ?e \* ?Q) ({At ?e} \ ?elm_f sr)" + proof - + have "pc ?e {At ?e}" by (simp add:pc_def sg_def) + (******************************************************************************) + moreover have "?Q (?elm_f sr)" + by (simp add:mm_def sg_def) + (*----------------------------------------------------------------------------*) + moreover + (******************************************************************************) + have "{At ?e} ## ?elm_f sr" by (simp add:set_ins_def) + (*----------------------------------------------------------------------------*) + ultimately show ?thesis by (auto intro!:sep_conjI simp:set_ins_def) + qed + moreover + (******************************************************************************) + from ia_disj.m_disj1 `?P sr` `s = {At i} \ sr` + have "?elm_f sr ## s'" by (auto simp:set_ins_def mm_def sg_def) + (*----------------------------------------------------------------------------*) + with at_fresh_s' + have fresh_atm: "{At ?e} \ ?elm_f sr ## s'" by (auto simp:set_ins_def) + moreover have eq_map: "\ elm. elm \ s' \ + (recse_map ks elm) = (recse_map (?ks_f sr ks) elm)" + proof - + fix elm + assume elm_in: "elm \ s'" + show "(recse_map ks elm) = (recse_map (?ks_f sr ks) elm)" + proof(cases elm) + (*******************************************************************) + case (M a' v') + from eq_s have "M a v \ s" by (auto simp:set_ins_def mm_def sg_def) + with elm_in ia_disj.m_disj1[OF this] M + have "a \ a'" by auto + thus ?thesis + apply (auto simp:M recse_map.simps pasrt_def list_ext_len) + apply (case_tac "a < length ks", auto simp:list_ext_len_eq list_ext_lt) + apply (case_tac "a' < Suc a", auto simp:list_ext_len_eq list_ext_lt) + by (metis (full_types) bot_nat_def + leI le_trans less_Suc_eq_le list_ext_lt_get list_ext_tail set_zero_def) + (*-----------------------------------------------------------------*) + qed auto + qed + ultimately show ?case + apply (rule_tac x = k in exI, rule_tac x = "{At ?e} \ ?elm_f sr" in exI) + apply (unfold IA_def, intro condI, assumption+) + apply (rule_tac x = "?ks_f sr ks" in pred_exI) + apply (rule_tac x = "?idx_f sr ks ia" in pred_exI) + apply (unfold fam_conj_disj_simp[OF fresh_atm]) + apply (auto simp:sep_conj_ac fam_conj_simps) + (***************************************************************************) + (* apply (unfold fam_conj_insert_simp [OF h_fresh1[OF at_fresh_sr]], simp) *) + (*-------------------------------------------------------------------------*) + apply (sep_cancel)+ + by (simp add: fam_conj_ext_eq[where I = "s'", OF eq_map]) +qed + +lemma tm_hoare_dec_fail: + assumes "mm a 0 sr" + shows + "\ fam_conj sr (recse_map ks) \* + st i \* ps 2 \* zero 0 \* zero 1 \* reps 2 ia ks \* fam_conj {ia<..} zero \ + i:[ (Dec a e) ]:j + \ fam_conj {M a 0} (recse_map (list_ext a ks[a := 0])) \* + st e \* + ps 2 \* + zero 0 \* + zero 1 \* + reps 2 (ia + int (reps_len (list_ext a ks)) - int (reps_len ks)) (list_ext a ks[a := 0]) \* + fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) <..} zero\" +proof - + from `mm a 0 sr` have eq_sr: "sr = {M a 0}" by (auto simp:mm_def sg_def) + { assume h: "a < length ks \ ks ! a = 0 \ length ks \ a" + from tm_hoare_dec_fail1[where u = 2, OF this] + have "\st i \* ps 2 \* zero 0 \* zero 1 \* reps 2 ia ks \* fam_conj {ia<..} zero\ + i :[ Dec a e ]: j + \st e \* ps 2 \* zero 0 \* zero 1 \* + reps 2 (ia + int (reps_len (list_ext a ks)) - int (reps_len ks)) (list_ext a ks[a := 0]) \* + fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks)<..} zero\" + by (simp) + } + thus ?thesis + apply (unfold eq_sr) + apply (simp add:fam_conj_simps list_ext_get_upd list_ext_len) + by (rule tm.pre_condI, blast) +qed + +lemma hoare_dec_fail: "IA. \ pc i ** mm a 0 \ + i:[ (Dec a e) ]:j + \ pc e ** mm a 0 \" + (is "IA. \ pc i ** ?P \ + i:[ ?code ?e]:j + \ pc ?e ** ?Q\") +proof(induct rule:tm.IHoareI) + case (IPre s' s r cnf) + let ?cnf = "(trset_of cnf)" + from IPre + have h: "(pc i \* ?P) s" "(s ## s')" "(IA (s + s') \* i :[ ?code(?e) ]: j \* r) ?cnf" + by (metis condD)+ + from h(1) obtain sr where + eq_s: "s = {At i} \ sr" "{At i} ## sr" "?P sr" + by (auto dest!:sep_conjD simp:set_ins_def pc_def sg_def) + hence "At i \ s" by auto + from h(3) obtain s1 s2 s3 + where hh: "?cnf = s1 + s2 + s3" + "s1 ## s2 \ s2 ## s3 \ s1 ## s3" + "IA (s + s') s1" + "(i :[ ?code ?e ]: j) s2" + "r s3" + apply (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac) + by (metis sep_add_commute sep_disj_addD1 sep_disj_addD2 sep_disj_commuteI) + interpret ia_disj: IA_disjoint s s' s1 cnf + proof + from `IA (s + s') s1` show "IA (s + s') s1" . + next + from `s ## s'` show "s ## s'" . + next + from hh(1) show "s1 \ ?cnf" by (auto simp:set_ins_def) + qed + from hh(1) have s1_belong: "s1 \ ?cnf" by (auto simp:set_ins_def) + from hh(3) + have "(EXS ks ia. + ps 2 \* + zero 0 \* + zero 1 \* + reps 2 ia ks \* + fam_conj {ia<..} zero \* + (st i \* fam_conj sr (recse_map ks)) \* fam_conj s' (recse_map ks)) + s1" + apply (unfold IA_def fam_conj_disj_simp[OF h(2)]) + apply (unfold eq_s) + by (insert eq_s(2), simp add: fam_conj_disj_simp[OF eq_s(2)] fam_conj_simps set_ins_def) + then obtain ks ia + where "((fam_conj sr (recse_map ks) \* st i \* ps 2 \* zero 0 \* zero 1 \* + reps 2 ia ks \* fam_conj {ia<..} zero) \* fam_conj s' (recse_map ks)) s1" + (is "(?PP \* ?QQ) s1") + by (unfold pred_ex_def, auto simp:sep_conj_ac) + then obtain ss1 ss2 where pres: + "s1 = ss1 + ss2" "ss1 ## ss2" + "?PP ss1" + "?QQ ss2" + by (auto elim!:sep_conjE intro!:sep_conjI) + from ia_disj.at_disj1 [OF `At i \ s`] + have at_fresh_s': "At ?e \ s'" . + have at_fresh_sr: "At ?e \ sr" + proof + assume at_in: "At ?e \ sr" + from h(3)[unfolded IA_def fam_conj_disj_simp[OF h(2)]] + have "TAt ?e \ trset_of cnf" + apply (elim EXS_elim1) + apply (unfold eq_s fam_conj_disj_simp[OF eq_s(2), unfolded set_ins_def] + fam_conj_elm_simp[OF at_in]) + apply (erule_tac sep_conjE, unfold set_ins_def)+ + by (auto simp:sep_conj_ac fam_conj_simps st_def sg_def tpn_set_def) + moreover from pres(1, 3) hh(1) have "TAt i \ trset_of cnf" + apply(erule_tac sep_conjE) + apply(erule_tac sep_conjE) + by (auto simp:st_def tpc_set_def sg_def set_ins_def) + ultimately have "i = ?e" + by (cases cnf, auto simp:tpn_set_def trset_of.simps) + from eq_s[unfolded this] at_in + show "False" by (auto simp:set_ins_def) + qed + from pres(3) and hh(2, 4, 5) pres(2, 4) + have pres1: "(?PP \* i :[ ?code(?e)]: j \* (r \* (fam_conj s' (recse_map ks)))) + (trset_of cnf)" + apply (unfold hh(1) pres(1)) + apply (rule sep_conjI[where x=ss1 and y = "ss2 + s2 + s3"], assumption+) + apply (rule sep_conjI[where x="s2" and y = "ss2 + s3"], assumption+) + apply (rule sep_conjI[where x="s3" and y = ss2], assumption+) + by (auto simp:set_ins_def) + (*****************************************************************************) + let ?ks_f = "\ sr ks. list_ext a ks[a:=0]" + let ?elm_f = "\ sr. {M a 0}" + let ?idx_f = "\ sr ks ia. (ia + int (reps_len (list_ext a ks)) - int (reps_len ks))" + (*----------------------------------------------------------------------------*) + (******************************************************************************) + from tm_hoare_dec_fail[OF `?P sr`, unfolded tm.Hoare_gen_def, rule_format, OF pres1] + obtain k where "((fam_conj (?elm_f sr) (recse_map (?ks_f sr ks)) \* + st ?e \* + ps 2 \* + zero 0 \* zero 1 \* reps 2 (?idx_f sr ks ia) (?ks_f sr ks) \* + fam_conj {?idx_f sr ks ia<..} zero) \* + i :[ ?code ?e ]: j \* r \* fam_conj s' (recse_map ks)) + (trset_of (sep_exec.run tstep (Suc k) cnf))" by blast + (*----------------------------------------------------------------------------*) + moreover have "(pc ?e \* ?Q) ({At ?e} \ ?elm_f sr)" + proof - + have "pc ?e {At ?e}" by (simp add:pc_def sg_def) + (******************************************************************************) + moreover have "?Q (?elm_f sr)" + by (simp add:mm_def sg_def) + (*----------------------------------------------------------------------------*) + moreover + (******************************************************************************) + have "{At ?e} ## ?elm_f sr" by (simp add:set_ins_def) + (*----------------------------------------------------------------------------*) + ultimately show ?thesis by (auto intro!:sep_conjI simp:set_ins_def) + qed + moreover + (******************************************************************************) + from ia_disj.m_disj1 `?P sr` `s = {At i} \ sr` + have "?elm_f sr ## s'" by (auto simp:set_ins_def mm_def sg_def) + (*----------------------------------------------------------------------------*) + with at_fresh_s' + have fresh_atm: "{At ?e} \ ?elm_f sr ## s'" by (auto simp:set_ins_def) + moreover have eq_map: "\ elm. elm \ s' \ + (recse_map ks elm) = (recse_map (?ks_f sr ks) elm)" + proof - + fix elm + assume elm_in: "elm \ s'" + show "(recse_map ks elm) = (recse_map (?ks_f sr ks) elm)" + proof(cases elm) + (*******************************************************************) + case (M a' v') + from eq_s have "M a 0 \ s" by (auto simp:set_ins_def mm_def sg_def) + with elm_in ia_disj.m_disj1[OF this] M + have "a \ a'" by auto + thus ?thesis + apply (auto simp:M recse_map.simps pasrt_def list_ext_len) + apply (case_tac "a < length ks", auto simp:list_ext_len_eq list_ext_lt) + apply (case_tac "a' < Suc a", auto simp:list_ext_len_eq list_ext_lt) + by (metis (full_types) bot_nat_def + leI le_trans less_Suc_eq_le list_ext_lt_get list_ext_tail set_zero_def) + (*-----------------------------------------------------------------*) + qed auto + qed + ultimately show ?case + apply (rule_tac x = k in exI, rule_tac x = "{At ?e} \ ?elm_f sr" in exI) + apply (unfold IA_def, intro condI, assumption+) + apply (rule_tac x = "?ks_f sr ks" in pred_exI) + apply (rule_tac x = "?idx_f sr ks ia" in pred_exI) + apply (unfold fam_conj_disj_simp[OF fresh_atm]) + apply (auto simp:sep_conj_ac fam_conj_simps) + (***************************************************************************) + (* apply (unfold fam_conj_insert_simp [OF h_fresh1[OF at_fresh_sr]], simp) *) + (*-------------------------------------------------------------------------*) + apply (sep_cancel)+ + by (simp add: fam_conj_ext_eq[where I = "s'", OF eq_map]) +qed + +lemma hoare_dec_fail_gen[step]: + assumes "v = 0" + shows + "IA. \ pc i ** mm a v \ + i:[ (Dec a e) ]:j + \ pc e ** mm a v \" + by (unfold assms, rule hoare_dec_fail) + + +lemma tm_hoare_dec_suc2: + assumes "mm a (Suc v) sr" + shows "\ fam_conj sr (recse_map ks) \* + st i \* ps 2 \* zero 0 \* zero 1 \* reps 2 ia ks \* fam_conj {ia<..} zero \ + i:[(Dec a e)]:j + \ fam_conj {M a v} (recse_map (list_ext a ks[a := v])) \* + st j \* + ps 2 \* + zero 0 \* + zero 1 \* + reps 2 (ia - 1) (list_ext a ks[a := v]) \* + fam_conj {ia - 1<..} zero\" +proof - + from `mm a (Suc v) sr` have eq_sr: "sr = {M a (Suc v)}" by (auto simp:mm_def sg_def) + thus ?thesis + apply (unfold eq_sr) + apply (simp add:fam_conj_simps list_ext_get_upd list_ext_len) + apply (rule tm.pre_condI) + by (drule tm_hoare_dec_suc1[where u = "2"], simp) +qed + +lemma hoare_dec_suc2: + "IA. \(pc i \* mm a (Suc v))\ + i :[ Dec a e ]: j + \pc j \* mm a v\" + (is "IA. \ pc i ** ?P \ + i:[ ?code ?e]:j + \ pc ?e ** ?Q\") +proof(induct rule:tm.IHoareI) + case (IPre s' s r cnf) + let ?cnf = "(trset_of cnf)" + from IPre + have h: "(pc i \* ?P) s" "(s ## s')" "(IA (s + s') \* i :[ ?code(?e) ]: j \* r) ?cnf" + by (metis condD)+ + from h(1) obtain sr where + eq_s: "s = {At i} \ sr" "{At i} ## sr" "?P sr" + by (auto dest!:sep_conjD simp:set_ins_def pc_def sg_def) + hence "At i \ s" by auto + from h(3) obtain s1 s2 s3 + where hh: "?cnf = s1 + s2 + s3" + "s1 ## s2 \ s2 ## s3 \ s1 ## s3" + "IA (s + s') s1" + "(i :[ ?code ?e ]: j) s2" + "r s3" + apply (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac) + by (metis sep_add_commute sep_disj_addD1 sep_disj_addD2 sep_disj_commuteI) + interpret ia_disj: IA_disjoint s s' s1 cnf + proof + from `IA (s + s') s1` show "IA (s + s') s1" . + next + from `s ## s'` show "s ## s'" . + next + from hh(1) show "s1 \ ?cnf" by (auto simp:set_ins_def) + qed + from hh(1) have s1_belong: "s1 \ ?cnf" by (auto simp:set_ins_def) + from hh(3) + have "(EXS ks ia. + ps 2 \* + zero 0 \* + zero 1 \* + reps 2 ia ks \* + fam_conj {ia<..} zero \* + (st i \* fam_conj sr (recse_map ks)) \* fam_conj s' (recse_map ks)) + s1" + apply (unfold IA_def fam_conj_disj_simp[OF h(2)]) + apply (unfold eq_s) + by (insert eq_s(2), simp add: fam_conj_disj_simp[OF eq_s(2)] fam_conj_simps set_ins_def) + then obtain ks ia + where "((fam_conj sr (recse_map ks) \* st i \* ps 2 \* zero 0 \* zero 1 \* + reps 2 ia ks \* fam_conj {ia<..} zero) \* fam_conj s' (recse_map ks)) s1" + (is "(?PP \* ?QQ) s1") + by (unfold pred_ex_def, auto simp:sep_conj_ac) + then obtain ss1 ss2 where pres: + "s1 = ss1 + ss2" "ss1 ## ss2" + "?PP ss1" + "?QQ ss2" + by (auto elim!:sep_conjE intro!:sep_conjI) + from ia_disj.at_disj1 [OF `At i \ s`] + have at_fresh_s': "At ?e \ s'" . + have at_fresh_sr: "At ?e \ sr" + proof + assume at_in: "At ?e \ sr" + from h(3)[unfolded IA_def fam_conj_disj_simp[OF h(2)]] + have "TAt ?e \ trset_of cnf" + apply (elim EXS_elim1) + apply (unfold eq_s fam_conj_disj_simp[OF eq_s(2), unfolded set_ins_def] + fam_conj_elm_simp[OF at_in]) + apply (erule_tac sep_conjE, unfold set_ins_def)+ + by (auto simp:sep_conj_ac fam_conj_simps st_def sg_def tpn_set_def) + moreover from pres(1, 3) hh(1) have "TAt i \ trset_of cnf" + apply(erule_tac sep_conjE) + apply(erule_tac sep_conjE) + by (auto simp:st_def tpc_set_def sg_def set_ins_def) + ultimately have "i = ?e" + by (cases cnf, auto simp:tpn_set_def trset_of.simps) + from eq_s[unfolded this] at_in + show "False" by (auto simp:set_ins_def) + qed + from pres(3) and hh(2, 4, 5) pres(2, 4) + have pres1: "(?PP \* i :[ ?code(?e)]: j \* (r \* (fam_conj s' (recse_map ks)))) + (trset_of cnf)" + apply (unfold hh(1) pres(1)) + apply (rule sep_conjI[where x=ss1 and y = "ss2 + s2 + s3"], assumption+) + apply (rule sep_conjI[where x="s2" and y = "ss2 + s3"], assumption+) + apply (rule sep_conjI[where x="s3" and y = ss2], assumption+) + by (auto simp:set_ins_def) + (*****************************************************************************) + let ?ks_f = "\ sr ks. list_ext a ks[a:=v]" + let ?elm_f = "\ sr. {M a v}" + let ?idx_f = "\ sr ks ia. ia - 1" + (*----------------------------------------------------------------------------*) + (******************************************************************************) + from tm_hoare_dec_suc2[OF `?P sr`, unfolded tm.Hoare_gen_def, rule_format, OF pres1] + obtain k where "((fam_conj (?elm_f sr) (recse_map (?ks_f sr ks)) \* + st ?e \* + ps 2 \* + zero 0 \* zero 1 \* reps 2 (?idx_f sr ks ia) (?ks_f sr ks) \* + fam_conj {?idx_f sr ks ia<..} zero) \* + i :[ ?code ?e ]: j \* r \* fam_conj s' (recse_map ks)) + (trset_of (sep_exec.run tstep (Suc k) cnf))" by blast + (*----------------------------------------------------------------------------*) + moreover have "(pc ?e \* ?Q) ({At ?e} \ ?elm_f sr)" + proof - + have "pc ?e {At ?e}" by (simp add:pc_def sg_def) + (******************************************************************************) + moreover have "?Q (?elm_f sr)" + by (simp add:mm_def sg_def) + (*----------------------------------------------------------------------------*) + moreover + (******************************************************************************) + have "{At ?e} ## ?elm_f sr" by (simp add:set_ins_def) + (*----------------------------------------------------------------------------*) + ultimately show ?thesis by (auto intro!:sep_conjI simp:set_ins_def) + qed + moreover + (******************************************************************************) + from ia_disj.m_disj1 `?P sr` `s = {At i} \ sr` + have "?elm_f sr ## s'" by (auto simp:set_ins_def mm_def sg_def) + (*----------------------------------------------------------------------------*) + with at_fresh_s' + have fresh_atm: "{At ?e} \ ?elm_f sr ## s'" by (auto simp:set_ins_def) + moreover have eq_map: "\ elm. elm \ s' \ + (recse_map ks elm) = (recse_map (?ks_f sr ks) elm)" + proof - + fix elm + assume elm_in: "elm \ s'" + show "(recse_map ks elm) = (recse_map (?ks_f sr ks) elm)" + proof(cases elm) + (*******************************************************************) + case (M a' v') + from eq_s have "M a (Suc v) \ s" by (auto simp:set_ins_def mm_def sg_def) + with elm_in ia_disj.m_disj1[OF this] M + have "a \ a'" by auto + thus ?thesis + apply (auto simp:M recse_map.simps pasrt_def list_ext_len) + apply (case_tac "a < length ks", auto simp:list_ext_len_eq list_ext_lt) + apply (case_tac "a' < Suc a", auto simp:list_ext_len_eq list_ext_lt) + by (metis (full_types) bot_nat_def + leI le_trans less_Suc_eq_le list_ext_lt_get list_ext_tail set_zero_def) + (*-----------------------------------------------------------------*) + qed auto + qed + ultimately show ?case + apply (rule_tac x = k in exI, rule_tac x = "{At ?e} \ ?elm_f sr" in exI) + apply (unfold IA_def, intro condI, assumption+) + apply (rule_tac x = "?ks_f sr ks" in pred_exI) + apply (rule_tac x = "?idx_f sr ks ia" in pred_exI) + apply (unfold fam_conj_disj_simp[OF fresh_atm]) + apply (auto simp:sep_conj_ac fam_conj_simps) + (***************************************************************************) + (* apply (unfold fam_conj_insert_simp [OF h_fresh1[OF at_fresh_sr]], simp) *) + (*-------------------------------------------------------------------------*) + apply (sep_cancel)+ + by (simp add: fam_conj_ext_eq[where I = "s'", OF eq_map]) +qed + +lemma hoare_dec_suc2_gen[step]: + assumes "v > 0" + shows + "IA. \pc i \* mm a v\ + i :[ Dec a e ]: j + \pc j \* mm a (v - 1)\" +proof - + from assms obtain v' where "v = Suc v'" + by (metis gr_implies_not0 nat.exhaust) + show ?thesis + apply (unfold `v = Suc v'`, simp) + by (rule hoare_dec_suc2) +qed + +definition [asmb]: "Goto e = jmp e" + +lemma hoare_jmp_reps2: + "\ st i \* ps u \* reps u v ks \* tm (v + 1) x \ + i:[(jmp e)]:j + \ st e \* ps u \* reps u v ks \* tm (v + 1) x \" +proof(cases "ks") + case Nil + thus ?thesis + by (simp add:sep_conj_cond reps.simps, intro tm.pre_condI, simp, hsteps) +next + case (Cons k ks') + thus ?thesis + proof(cases "ks' = []") + case True with Cons + show ?thesis + apply(simp add:sep_conj_cond reps.simps, intro tm.pre_condI, simp add:ones_simps) + by (hgoto hoare_jmp[where p = u]) + next + case False + show ?thesis + apply (unfold `ks = k#ks'` reps_simp3[OF False], simp add:ones_simps) + by (hgoto hoare_jmp[where p = u]) + qed +qed + +lemma tm_hoare_goto_pre: (* ccc *) + assumes "() sr" + shows "\ fam_conj sr (recse_map ks) \* + st i \* ps 2 \* zero 0 \* zero 1 \* reps 2 ia ks \* fam_conj {ia<..} zero \ + i:[(Goto e)]:j + \ fam_conj {} (recse_map ks) \* + st e \* ps 2 \* zero 0 \* zero 1 \* reps 2 ia ks \* fam_conj {ia<..} zero \" + apply (unfold Goto_def) + apply (subst (1 2) fam_conj_interv_simp) + apply (unfold zero_def) + apply (hstep hoare_jmp_reps2) + apply (simp add:sep_conj_ac) + my_block + from assms have "sr = {}" + by (simp add:pasrt_def set_ins_def) + my_block_end + by (unfold this, sep_cancel+) + +lemma hoare_goto_pre: + "IA. \ pc i \* \ + i:[ (Goto e) ]:j + \ pc e \* \" + (is "IA. \ pc i ** ?P \ + i:[ ?code ?e]:j + \ pc ?e ** ?Q\") +proof(induct rule:tm.IHoareI) + case (IPre s' s r cnf) + let ?cnf = "(trset_of cnf)" + from IPre + have h: "(pc i \* ?P) s" "(s ## s')" "(IA (s + s') \* i :[ ?code(?e) ]: j \* r) ?cnf" + by (metis condD)+ + from h(1) obtain sr where + eq_s: "s = {At i} \ sr" "{At i} ## sr" "?P sr" + by (auto dest!:sep_conjD simp:set_ins_def pc_def sg_def pasrt_def) + hence "At i \ s" by auto + from h(3) obtain s1 s2 s3 + where hh: "?cnf = s1 + s2 + s3" + "s1 ## s2 \ s2 ## s3 \ s1 ## s3" + "IA (s + s') s1" + "(i :[ ?code ?e ]: j) s2" + "r s3" + apply (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac) + by (metis sep_add_commute sep_disj_addD1 sep_disj_addD2 sep_disj_commuteI) + interpret ia_disj: IA_disjoint s s' s1 cnf + proof + from `IA (s + s') s1` show "IA (s + s') s1" . + next + from `s ## s'` show "s ## s'" . + next + from hh(1) show "s1 \ ?cnf" by (auto simp:set_ins_def) + qed + from hh(1) have s1_belong: "s1 \ ?cnf" by (auto simp:set_ins_def) + from hh(3) + have "(EXS ks ia. + ps 2 \* + zero 0 \* + zero 1 \* + reps 2 ia ks \* + fam_conj {ia<..} zero \* + (st i \* fam_conj sr (recse_map ks)) \* fam_conj s' (recse_map ks)) + s1" + apply (unfold IA_def fam_conj_disj_simp[OF h(2)]) + apply (unfold eq_s) + by (insert eq_s(2), simp add: fam_conj_disj_simp[OF eq_s(2)] fam_conj_simps set_ins_def) + then obtain ks ia + where "((fam_conj sr (recse_map ks) \* st i \* ps 2 \* zero 0 \* zero 1 \* + reps 2 ia ks \* fam_conj {ia<..} zero) \* fam_conj s' (recse_map ks)) s1" + (is "(?PP \* ?QQ) s1") + by (unfold pred_ex_def, auto simp:sep_conj_ac) + then obtain ss1 ss2 where pres: + "s1 = ss1 + ss2" "ss1 ## ss2" + "?PP ss1" + "?QQ ss2" + by (auto elim!:sep_conjE intro!:sep_conjI) + from ia_disj.at_disj1 [OF `At i \ s`] + have at_fresh_s': "At ?e \ s'" . + have at_fresh_sr: "At ?e \ sr" + proof + assume at_in: "At ?e \ sr" + from h(3)[unfolded IA_def fam_conj_disj_simp[OF h(2)]] + have "TAt ?e \ trset_of cnf" + apply (elim EXS_elim1) + apply (unfold eq_s fam_conj_disj_simp[OF eq_s(2), unfolded set_ins_def] + fam_conj_elm_simp[OF at_in]) + apply (erule_tac sep_conjE, unfold set_ins_def)+ + by (auto simp:sep_conj_ac fam_conj_simps st_def sg_def tpn_set_def) + moreover from pres(1, 3) hh(1) have "TAt i \ trset_of cnf" + apply(erule_tac sep_conjE) + apply(erule_tac sep_conjE) + by (auto simp:st_def tpc_set_def sg_def set_ins_def) + ultimately have "i = ?e" + by (cases cnf, auto simp:tpn_set_def trset_of.simps) + from eq_s[unfolded this] at_in + show "False" by (auto simp:set_ins_def) + qed + from pres(3) and hh(2, 4, 5) pres(2, 4) + have pres1: "(?PP \* i :[ ?code(?e)]: j \* (r \* (fam_conj s' (recse_map ks)))) + (trset_of cnf)" + apply (unfold hh(1) pres(1)) + apply (rule sep_conjI[where x=ss1 and y = "ss2 + s2 + s3"], assumption+) + apply (rule sep_conjI[where x="s2" and y = "ss2 + s3"], assumption+) + apply (rule sep_conjI[where x="s3" and y = ss2], assumption+) + by (auto simp:set_ins_def) + (*****************************************************************************) + let ?ks_f = "\ sr ks. ks" + let ?elm_f = "\ sr. {}" + let ?idx_f = "\ sr ks ia. ia" + (*----------------------------------------------------------------------------*) + (******************************************************************************) + from tm_hoare_goto_pre[OF `?P sr`, unfolded tm.Hoare_gen_def, rule_format, OF pres1] + obtain k where "((fam_conj (?elm_f sr) (recse_map (?ks_f sr ks)) \* + st ?e \* + ps 2 \* + zero 0 \* zero 1 \* reps 2 (?idx_f sr ks ia) (?ks_f sr ks) \* + fam_conj {?idx_f sr ks ia<..} zero) \* + i :[ ?code ?e ]: j \* r \* fam_conj s' (recse_map ks)) + (trset_of (sep_exec.run tstep (Suc k) cnf))" by blast + (*----------------------------------------------------------------------------*) + moreover have "(pc ?e \* ?Q) ({At ?e} \ ?elm_f sr)" + proof - + have "pc ?e {At ?e}" by (simp add:pc_def sg_def) + (******************************************************************************) + moreover have "?Q (?elm_f sr)" + by (simp add:pasrt_def set_ins_def) + (*----------------------------------------------------------------------------*) + moreover + (******************************************************************************) + have "{At ?e} ## ?elm_f sr" by (simp add:set_ins_def) + (*----------------------------------------------------------------------------*) + ultimately show ?thesis by (auto intro!:sep_conjI simp:set_ins_def) + qed + moreover + (******************************************************************************) + from ia_disj.m_disj1 `?P sr` `s = {At i} \ sr` + have "?elm_f sr ## s'" by (auto simp:set_ins_def mm_def sg_def) + (*----------------------------------------------------------------------------*) + with at_fresh_s' + have fresh_atm: "{At ?e} \ ?elm_f sr ## s'" by (auto simp:set_ins_def) + moreover have eq_map: "\ elm. elm \ s' \ + (recse_map ks elm) = (recse_map (?ks_f sr ks) elm)" + proof - + fix elm + assume elm_in: "elm \ s'" + show "(recse_map ks elm) = (recse_map (?ks_f sr ks) elm)" + by simp + qed + ultimately show ?case + apply (rule_tac x = k in exI, rule_tac x = "{At ?e} \ ?elm_f sr" in exI) + apply (unfold IA_def, intro condI, assumption+) + apply (rule_tac x = "?ks_f sr ks" in pred_exI) + apply (rule_tac x = "?idx_f sr ks ia" in pred_exI) + apply (unfold fam_conj_disj_simp[OF fresh_atm]) + by (auto simp:sep_conj_ac fam_conj_simps) +qed + +lemma hoare_goto[step]: "IA. \ pc i \ + i:[ (Goto e) ]:j + \ pc e \" +proof(rule tm.I_hoare_adjust [OF hoare_goto_pre]) + fix s assume "pc i s" thus "(pc i \* ) s" + by (metis cond_true_eq2) +next + fix s assume "(pc e \* ) s" thus "pc e s" + by (metis cond_true_eq2) +qed + +lemma I_hoare_sequence: + assumes h1: "\ i j. I. \pc i ** p\ i:[c1]:j \pc j ** q\" + and h2: "\ j k. I. \pc j ** q\ j:[c2]:k \pc k ** r\" + shows "I. \pc i ** p\ i:[(c1 ; c2)]:k \pc k ** r\" +proof(unfold tassemble_to.simps, intro tm.I_code_exI) + fix j' + show "I.\pc i \* p\ i :[ c1 ]: j' \* j' :[ c2 ]: k \pc k \* r\" + proof(rule tm.I_sequencing) + from tm.I_code_extension[OF h1 [of i j'], of" j' :[ c2 ]: k"] + show "I.\pc i \* p\ i :[ c1 ]: j' \* j' :[ c2 ]: k \pc j' \* q\" . + next + from tm.I_code_extension[OF h2 [of j' k], of" i :[ c1 ]: j'"] + show "I.\pc j' \* q\ i :[ c1 ]: j' \* j' :[ c2 ]: k \pc k \* r\" + by (auto simp:sep_conj_ac) + qed +qed + +lemma I_hoare_seq1: + assumes h1: "\j'. I. \pc i ** p\ i:[c1]:j' \pc j' ** q\" + and h2: "\j' . I. \pc j' ** q\ j':[c2]:k \pc k' ** r\" + shows "I. \pc i ** p\ i:[(c1 ; c2)]:k \pc k' ** r\" +proof(unfold tassemble_to.simps, intro tm.I_code_exI) + fix j' + show "I.\pc i \* p\ i :[ c1 ]: j' \* j' :[ c2 ]: k \pc k' \* r\" + proof(rule tm.I_sequencing) + from tm.I_code_extension[OF h1 [of j'], of "j' :[ c2 ]: k "] + show "I.\pc i \* p\ i :[ c1 ]: j' \* j' :[ c2 ]: k \pc j' \* q\" . + next + from tm.I_code_extension[OF h2 [of j'], of" i :[ c1 ]: j'"] + show "I.\pc j' \* q\ i :[ c1 ]: j' \* j' :[ c2 ]: k \pc k' \* r\" + by (auto simp:sep_conj_ac) + qed +qed + +lemma t_hoare_local1: + "(\l. \p\ i :[ c l ]: j \q\) \ + \p\ i:[TLocal c]:j \q\" +by (unfold tassemble_to.simps, rule tm.code_exI, auto) + +lemma I_hoare_local: + assumes h: "(\l. I.\pc i \* p\ i :[ c l ]: j \pc k \* q\)" + shows "I. \pc i ** p\ i:[TLocal c]:j \pc k ** q\" +proof(unfold tassemble_to.simps, rule tm.I_code_exI) + fix l + from h[of l] + show " I.\pc i \* p\ i :[ c l ]: j \pc k \* q\" . +qed + +lemma t_hoare_label1: + "(\l. l = i \ \p\ l :[ c (l::tstate) ]: j \q\) \ + \p \ + i:[(TLabel l; c l)]:j + \q\" +by +(unfold tassemble_to.simps, intro tm.code_exI tm.code_condI, clarify, case_tac l, auto) + +lemma I_hoare_label: + assumes h:"\l. l = i \ I. \pc l \* p\ l :[ c (l::tstate) ]: j \pc k \* q\" + shows "I. \pc i \* p \ + i:[(TLabel l; c l)]:j + \pc k \* q\" +proof(unfold tm.IHoare_def, default) + fix s' + show " \EXS s. <(pc i \* p) s> \* <(s ## s')> \* I (s + s')\ i :[ (TLabel l ; c l) ]: j + \EXS s. <(pc k \* q) s> \* <(s ## s')> \* I (s + s')\" + proof(rule t_hoare_label1) + fix l assume "l = i" + from h[OF this, unfolded tm.IHoare_def] + show "\EXS s. <(pc i \* p) s> \* <(s ## s')> \* I (s + s')\ l :[ c l ]: j + \EXS s. <(pc k \* q) s> \* <(s ## s')> \* I (s + s')\" + by (simp add:`l = i`) + qed +qed + +lemma I_hoare_label_last: + assumes h1: "t_last_cmd c = Some (TLabel l)" + and h2: "l = j \ I. \p\ i:[t_blast_cmd c]:j \q\" + shows "I. \p\ i:[c]:j \q\" +proof(unfold tm.IHoare_def, default) + fix s' + show "\EXS s.

\* <(s ## s')> \* I (s + s')\ i :[ c ]: j + \EXS s. \* <(s ## s')> \* I (s + s')\" + proof(rule t_hoare_label_last[OF h1]) + assume "l = j" + from h2[OF this, unfolded tm.IHoare_def] + show "\EXS s.

\* <(s ## s')> \* I (s + s')\ i :[ t_blast_cmd c ]: j + \EXS s. \* <(s ## s')> \* I (s + s')\" + by fast + qed +qed + +lemma I_hoare_seq2: + assumes h: "\j'. I. \pc i ** p\ i:[c1]:j' \pc k' \* r\" + shows "I. \pc i ** p\ i:[(c1 ; c2)]:j \pc k' ** r\" + apply (unfold tassemble_to.simps, intro tm.I_code_exI) + apply (unfold tm.IHoare_def, default) + apply (rule tm.code_extension) + by (rule h[unfolded tm.IHoare_def, rule_format]) + +lemma IA_pre_stren: + assumes h1: "IA. \p\ c \q\" + and h2: "\s. r s \ p s" + shows "IA. \r\ c \q\" + by (rule tm.I_pre_stren[OF assms], simp) + +lemma IA_post_weaken: + assumes h1: "IA. \p\ c \q\" + and h2: "\ s. q s \ r s" + shows "IA. \p\ c \r\" + by (rule tm.I_post_weaken[OF assms], simp) + +section {* Making triple processor for IA *} + +ML {* (* Functions specific to Hoare triple: IA {P} c {Q} *) + fun get_pre ctxt t = + let val pat = term_of @{cpat "IA. \?P\ ?c \?Q\"} + val env = match ctxt pat t + in inst env (term_of @{cpat "?P::aresource set \ bool"}) end + + fun can_process ctxt t = ((get_pre ctxt t; true) handle _ => false) + + fun get_post ctxt t = + let val pat = term_of @{cpat "IA. \?P\ ?c \?Q\"} + val env = match ctxt pat t + in inst env (term_of @{cpat "?Q::aresource set \ bool"}) end; + + fun get_mid ctxt t = + let val pat = term_of @{cpat "IA. \?P\ ?c \?Q\"} + val env = match ctxt pat t + in inst env (term_of @{cpat "?c::tresource set \ bool"}) end; + + fun is_pc_term (Const (@{const_name pc}, _) $ _) = true + | is_pc_term _ = false + + fun mk_pc_term x = + Const (@{const_name pc}, @{typ "nat \ aresource set \ bool"}) $ Free (x, @{typ "nat"}) + + val sconj_term = term_of @{cterm "sep_conj::assert \ assert \ assert"} + + val abc_triple = {binding = @{binding "abc_triple"}, + can_process = can_process, + get_pre = get_pre, + get_mid = get_mid, + get_post = get_post, + is_pc_term = is_pc_term, + mk_pc_term = mk_pc_term, + sconj_term = sconj_term, + sep_conj_ac_tac = sep_conj_ac_tac, + hoare_seq1 = @{thm I_hoare_seq1}, + hoare_seq2 = @{thm I_hoare_seq2}, + pre_stren = @{thm IA_pre_stren}, + post_weaken = @{thm IA_post_weaken}, + frame_rule = @{thm tm.I_frame_rule} + }:HoareTriple + + val _ = (HoareTriples_get ()) |> (fn orig => HoareTriples_store (abc_triple::orig)) +*} + +section {* Example proofs *} + +definition [asmb]: "clear a = (TL start exit. TLabel start; Dec a exit; Goto start; TLabel exit)" + +lemma hoare_clear[step]: + "IA. \pc i ** mm a v\ + i:[clear a]:j + \pc j ** mm a 0\" +proof(unfold clear_def, intro I_hoare_local I_hoare_label, simp, + rule I_hoare_label_last, simp+, prune) + show "IA.\pc i \* mm a v\ i :[ (Dec a j ; Goto i) ]: j \pc j \* mm a 0\" + proof(induct v) + case 0 + show ?case + by hgoto + next + case (Suc v) + show ?case + apply (rule_tac Q = "pc i \* mm a v" in tm.I_sequencing) + by hsteps + qed +qed + +definition [asmb]: + "dup a b c = + (TL start exit. TLabel start; Dec a exit; Inc b; Inc c; Goto start; TLabel exit)" + +lemma hoare_dup[step]: + "IA. \pc i ** mm a va ** mm b vb ** mm c vc \ + i:[dup a b c]:j + \pc j ** mm a 0 ** mm b (va + vb) ** mm c (va + vc)\" +proof(unfold dup_def, intro I_hoare_local I_hoare_label, clarsimp, + rule I_hoare_label_last, simp+, prune) + show "IA. \pc i \* mm a va \* mm b vb \* mm c vc\ + i :[ (Dec a j ; Inc b ; Inc c ; Goto i) ]: j + \pc j \* mm a 0 \* mm b (va + vb) \* mm c (va + vc)\" + proof(induct va arbitrary: vb vc) + case (0 vb vc) + show ?case + by hgoto + next + case (Suc va vb vc) + show ?case + apply (rule_tac Q = "pc i \* mm a va \* mm b (Suc vb) \* mm c (Suc vc)" in tm.I_sequencing) + by (hsteps Suc) + qed +qed + +definition [asmb]: + "clear_add a b = + (TL start exit. TLabel start; Dec a exit; Inc b; Goto start; TLabel exit)" + +lemma hoare_clear_add[step]: + "IA. \pc i ** mm a va ** mm b vb \ + i:[clear_add a b]:j + \pc j ** mm a 0 ** mm b (va + vb)\" +proof(unfold clear_add_def, intro I_hoare_local I_hoare_label, clarsimp, + rule I_hoare_label_last, simp+, prune) + show "IA. \pc i \* mm a va \* mm b vb\ + i :[ (Dec a j ; Inc b ; Goto i) ]: j + \pc j \* mm a 0 \* mm b (va + vb)\" + proof(induct va arbitrary: vb) + case 0 + show ?case + by hgoto + next + case (Suc va vb) + show ?case + apply (rule_tac Q = "pc i \* mm a va \* mm b (Suc vb)" in tm.I_sequencing) + by (hsteps Suc) + qed +qed + +definition [asmb]: + "copy_to a b c = clear b; clear c; dup a b c; clear_add c a" + +lemma hoare_copy_to[step]: + "IA. \pc i ** mm a va ** mm b vb ** mm c vc \ + i:[copy_to a b c]:j + \pc j ** mm a va ** mm b va ** mm c 0\" + by (unfold copy_to_def, hsteps) + +definition [asmb]: + "preserve_add a b c = clear c; dup a b c; clear_add c a" + +lemma hoare_preserve_add[step]: + "IA. \pc i ** mm a va ** mm b vb ** mm c vc \ + i:[preserve_add a b c]:j + \pc j ** mm a va ** mm b (va + vb) ** mm c 0\" + by (unfold preserve_add_def, hsteps) + +definition [asmb]: + "mult a b c t1 t2 = + clear c; + copy_to a t2 t1; + (TL start exit. + TLabel start; + Dec a exit; + preserve_add b c t1; + Goto start; + TLabel exit + ); + clear_add t2 a" + +lemma hoare_mult[step]: + "IA. \pc i ** mm a va ** mm b vb ** mm c vc ** mm t1 vt1 ** mm t2 vt2 \ + i:[mult a b c t1 t2]:j + \pc j ** mm a va ** mm b vb ** mm c (va * vb) ** mm t1 0 ** mm t2 0 \" + apply (unfold mult_def, hsteps) + apply (rule_tac q = "mm a 0 \* mm b vb \* mm c (va * vb) \* mm t1 0 \* mm t2 va" in I_hoare_seq1) + apply (intro I_hoare_local I_hoare_label, clarify, + rule I_hoare_label_last, simp+, clarify, prune) + my_block + fix i j vc + have "IA. \pc i \* mm a va \* mm t1 0 \* mm c vc \* mm b vb\ + i :[ (Dec a j ; preserve_add b c t1 ; Goto i) ]: j + \pc j \* mm a 0 \* mm b vb \* mm c (va * vb + vc) \* mm t1 0 \" + proof(induct va arbitrary:vc) + case (0 vc) + show ?case + by hgoto + next + case (Suc va vc) + show ?case + apply (rule_tac Q = "pc i \* mm a va \* mm t1 0 \* mm c (vb + vc) \* mm b vb" + in tm.I_sequencing) + apply (hsteps Suc) + by (sep_cancel+, simp, smt) + qed + my_block_end + by (hsteps this) + +end + diff -r 77daf1b85cf0 -r a5f5b9336007 thys2/Hoare_gen.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/Hoare_gen.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,601 @@ +header {* + Generic Separation Logic +*} + +theory Hoare_gen +imports Main + "../progtut/Tactical" "../Separation_Algebra/Sep_Tactics" +begin + +ML_file "../Separation_Algebra/sep_tactics.ML" + +definition pasrt :: "bool \ (('a::sep_algebra) \ bool)" ("<_>" [72] 71) +where "pasrt b = (\ s . s = 0 \ b)" + +lemma sep_conj_cond1: "(p \* \* q) = ( \* p \* q)" + by(simp add: sep_conj_ac) + +lemma sep_conj_cond2: "(p \* ) = ( \* p)" + by(simp add: sep_conj_ac) + +lemma sep_conj_cond3: "(( \* p) \* r) = ( \* p \* r)" + by (metis sep.mult_assoc) + +lemmas sep_conj_cond = sep_conj_cond1 sep_conj_cond2 sep_conj_cond3 + +lemma cond_true_eq[simp]: " = \" + by(unfold sep_empty_def pasrt_def, auto) + +lemma cond_true_eq1: "( \* p) = p" + by(simp) + +lemma false_simp [simp]: " = sep_false" (* move *) + by (simp add:pasrt_def) + +lemma cond_true_eq2: "(p \* ) = p" + by simp + +lemma condD: "( ** r) s \ b \ r s" +by (unfold sep_conj_def pasrt_def, auto) + +locale sep_exec = + fixes step :: "'conf \ 'conf" + and recse:: "'conf \ 'a::sep_algebra" +begin + +definition "run n = step ^^ n" + +lemma run_add: "run (n1 + n2) s = run n1 (run n2 s)" + apply (unfold run_def) + by (metis funpow_add o_apply) + +definition + Hoare_gen :: "('a \ bool) \ ('a \ bool) \ ('a \ bool) \ bool" + ("(\(1_)\ / (_)/ \(1_)\)" 50) +where + "\ p \ c \ q \ \ + (\ s r. (p**c**r) (recse s) \ (\ k. ((q ** c ** r) (recse (run (Suc k) s)))))" + +lemma HoareI [case_names Pre]: + assumes h: "\ r s. (p**c**r) (recse s) \ (\ k. ((q ** c ** r) (recse (run (Suc k) s))))" + shows "\ p \ c \ q \" + using h + by (unfold Hoare_gen_def, auto) + +lemma frame_rule: + assumes h: "\ p \ c \ q \" + shows "\ p ** r \ c \ q ** r \" +proof(induct rule: HoareI) + case (Pre r' s') + hence "(p \* c \* r \* r') (recse s')" by (auto simp:sep_conj_ac) + from h[unfolded Hoare_gen_def, rule_format, OF this] + show ?case + by (metis sep_conj_assoc sep_conj_left_commute) +qed + +lemma sequencing: + assumes h1: "\p\ c \q\" + and h2: "\q\ c \r\" + shows "\p\ c \r\" +proof(induct rule:HoareI) + case (Pre r' s') + from h1[unfolded Hoare_gen_def, rule_format, OF Pre] + obtain k1 where "(q \* c \* r') (recse (run (Suc k1) s'))" by auto + from h2[unfolded Hoare_gen_def, rule_format, OF this] + obtain k2 where "(r \* c \* r') (recse (run (Suc k2) (run (Suc k1) s')))" by auto + thus ?case + apply (rule_tac x = "Suc (k1 + k2)" in exI) + by (metis add_Suc_right nat_add_commute sep_exec.run_add) +qed + +lemma pre_stren: + assumes h1: "\p\ c \q\" + and h2: "\s. r s \ p s" + shows "\r\ c \q\" +proof(induct rule:HoareI) + case (Pre r' s') + with h2 + have "(p \* c \* r') (recse s')" + by (metis sep_conj_impl1) + from h1[unfolded Hoare_gen_def, rule_format, OF this] + show ?case . +qed + +lemma post_weaken: + assumes h1: "\p\ c \q\" + and h2: "\ s. q s \ r s" + shows "\p\ c \r\" +proof(induct rule:HoareI) + case (Pre r' s') + from h1[unfolded Hoare_gen_def, rule_format, OF this] + obtain k where "(q \* c \* r') (recse (run (Suc k) s'))" by blast + with h2 + show ?case + by (metis sep_conj_impl1) +qed + +lemma hoare_adjust: + assumes h1: "\p1\ c \q1\" + and h2: "\s. p s \ p1 s" + and h3: "\s. q1 s \ q s" + shows "\p\ c \q\" + using h1 h2 h3 post_weaken pre_stren + by (metis) + +lemma code_exI: + assumes h: "\ k. \p\ c(k) \q\" + shows "\p\ EXS k. c(k) \q\" +proof(unfold pred_ex_def, induct rule:HoareI) + case (Pre r' s') + then obtain k where "(p \* (\ s. c k s) \* r') (recse s')" + by (auto elim!:sep_conjE intro!:sep_conjI) + from h[unfolded Hoare_gen_def, rule_format, OF this] + show ?case + by (auto elim!:sep_conjE intro!:sep_conjI) +qed + +lemma code_extension: + assumes h: "\ p \ c \ q \" + shows "\ p \ c ** e \ q \" +proof(induct rule:HoareI) + case (Pre r' s') + hence "(p \* c \* e \* r') (recse s')" by (auto simp:sep_conj_ac) + from h[unfolded Hoare_gen_def, rule_format, OF this] + show ?case + by (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac) +qed + +lemma code_extension1: + assumes h: "\ p \ c \ q \" + shows "\ p \ e ** c \ q \" + by (metis code_extension h sep.mult_commute) + +lemma composition: + assumes h1: "\p\ c1 \q\" + and h2: "\q\ c2 \r\" + shows "\p\ c1 ** c2 \r\" +proof(induct rule:HoareI) + case (Pre r' s') + hence "(p \* c1 \* c2 \* r') (recse s')" by (auto simp:sep_conj_ac) + from h1[unfolded Hoare_gen_def, rule_format, OF this] + obtain k1 where "(q \* c2 \* c1 \* r') (recse (run (Suc k1) s'))" + by (auto simp:sep_conj_ac) + from h2[unfolded Hoare_gen_def, rule_format, OF this, folded run_add] + show ?case + by (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac) +qed + + +definition + IHoare :: "('b::sep_algebra \ 'a \ bool) \ + ('b \ bool) \ ('a \ bool) \ ('b \ bool) \ bool" + ("(1_).(\(1_)\ / (_)/ \(1_)\)" 50) +where + "I. \P\ c \Q\ = (\s'. \ EXS s.

\* <(s ## s')> \* I(s + s')\ c + \ EXS s. \* <(s ## s')> \* I(s + s')\)" + +lemma IHoareI [case_names IPre]: + assumes h: "\s' s r cnf . (

\* <(s ## s')> \* I(s + s') \* c \* r) (recse cnf) \ + (\k t. ( \* <(t ## s')> \* I(t + s') \* c \* r) (recse (run (Suc k) cnf)))" + shows "I. \P\ c \Q\" + unfolding IHoare_def +proof + fix s' + show " \EXS s.

\* <(s ## s')> \* I (s + s')\ c + \EXS s. \* <(s ## s')> \* I (s + s')\" + proof(unfold pred_ex_def, induct rule:HoareI) + case (Pre r s) + then obtain sa where "(

\* <(sa ## s')> \* I (sa + s') \* c \* r) (recse s)" + by (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac) + hence " (\k t. ( \* <(t##s')> \* I(t + s') \* c \* r) (recse (run (Suc k) s)))" + by (rule h) + then obtain k t where h2: "( \* <(t ## s')> \* I(t + s') \* c \* r) (recse (run (Suc k) s))" + by auto + thus "\k. ((\s. \sa. ( \* <(sa ## s')> \* I (sa + s')) s) \* c \* r) (recse (run (Suc k) s))" + apply (rule_tac x = k in exI) + by (auto intro!:sep_conjI elim!:sep_conjE simp:sep_conj_ac) + qed + qed + +lemma I_frame_rule: + assumes h: "I. \P\ c \Q\" + shows "I. \P \* R\ c \Q \* R\" +proof(induct rule:IHoareI) + case (IPre s' s r cnf) + hence "((<(P \* R) s> \* <(s##s')> \* I (s + s')) \* c \* r) (recse cnf)" + by (auto simp:sep_conj_ac) + then obtain s1 s2 + where h1: "((

\* <((s1 + s2) ## s')> \*I (s1 + s2 + s')) \* c \* r) (recse cnf)" + "s1 ## s2" "s1 + s2 ## s'" "P s1" "R s2" + by (unfold pasrt_def, auto elim!:sep_conjE intro!:sep_conjI) + hence "((EXS s.

\* <(s ## s2 +s')> \*I (s + (s2 + s'))) \* c \* r) (recse cnf)" + apply (sep_cancel, unfold pred_ex_def, auto intro!:sep_conjI sep_disj_addI3 elim!:sep_conjE) + apply (rule_tac x = s1 in exI, unfold pasrt_def, + auto intro!:sep_conjI sep_disj_addI3 elim!:sep_conjE simp:sep_conj_ac) + by (metis sep_add_assoc sep_add_disjD) + from h[unfolded IHoare_def Hoare_gen_def, rule_format, OF this] + obtain k s where + "(( \* <(s ## s2 + s')> \* I (s + (s2 + s'))) \* c \* r) (recse (run (Suc k) cnf))" + by (unfold pasrt_def pred_ex_def, auto elim!:sep_conjE intro!:sep_conjI) + thus ?case + proof(rule_tac x = k in exI, rule_tac x = "s + s2" in exI, sep_cancel+) + fix h ha + assume hh: "( \* <(s ## s2 + s')> \* I (s + (s2 + s'))) ha" + show " (<(Q \* R) (s + s2)> \* <(s + s2 ## s')> \* I (s + s2 + s')) ha" + proof - + from hh have h0: "s ## s2 + s'" + by (metis pasrt_def sep_conjD) + with h1(2, 3) + have h2: "s + s2 ## s'" + by (metis sep_add_disjD sep_disj_addI1) + moreover from h1(2, 3) h2 have h3: "(s + (s2 + s')) = (s + s2 + s')" + by (metis `s ## s2 + s'` sep_add_assoc sep_add_disjD sep_disj_addD1) + moreover from hh have "Q s" by (metis pasrt_def sep_conjD) + moreover from h0 h1(2) h1(3) have "s ## s2" + by (metis sep_add_disjD sep_disj_addD) + moreover note h1(5) + ultimately show ?thesis + by (smt h0 hh sep_conjI) + qed + qed +qed + +lemma I_sequencing: + assumes h1: "I. \P\ c \Q\" + and h2: "I. \Q\ c \R\" + shows "I. \P\ c \R\" + using h1 h2 sequencing + by (smt IHoare_def) + +lemma I_pre_stren: + assumes h1: "I. \P\ c \Q\" + and h2: "\s. R s \ P s" + shows "I. \R\ c \Q\" +proof(unfold IHoare_def, default) + fix s' + show "\EXS s. \* <(s ## s')> \* I (s + s')\ c + \EXS s. \* <(s ## s')> \* I (s + s')\" + proof(rule pre_stren) + from h1[unfolded IHoare_def, rule_format, of s'] + show "\EXS s.

\* <(s ## s')> \* I (s + s')\ c + \EXS s. \* <(s ## s')> \* I (s + s')\" . + next + fix s + show "(EXS s. \* <(s ## s')> \* I (s + s')) s \ + (EXS s.

\* <(s ## s')> \* I (s + s')) s" + apply (unfold pred_ex_def, clarify) + apply (rule_tac x = sa in exI, sep_cancel+) + by (insert h2, auto simp:pasrt_def) + qed +qed + +lemma I_post_weaken: + assumes h1: "I. \P\ c \Q\" + and h2: "\ s. Q s \ R s" + shows "I. \P\ c \R\" +proof(unfold IHoare_def, default) + fix s' + show "\EXS s.

\* <(s ## s')> \* I (s + s')\ c + \EXS s. \* <(s ## s')> \* I (s + s')\" + proof(rule post_weaken) + from h1[unfolded IHoare_def, rule_format, of s'] + show "\EXS s.

\* <(s ## s')> \* I (s + s')\ c + \EXS s. \* <(s ## s')> \* I (s + s')\" . + next + fix s + show "(EXS s. \* <(s ## s')> \* I (s + s')) s \ + (EXS s. \* <(s ## s')> \* I (s + s')) s" + apply (unfold pred_ex_def, clarify) + apply (rule_tac x = sa in exI, sep_cancel+) + by (insert h2, auto simp:pasrt_def) + qed +qed + +lemma I_hoare_adjust: + assumes h1: "I. \P1\ c \Q1\" + and h2: "\s. P s \ P1 s" + and h3: "\s. Q1 s \ Q s" + shows "I. \P\ c \Q\" + using h1 h2 h3 I_post_weaken I_pre_stren + by (metis) + +lemma I_code_exI: + assumes h: "\ k. I. \P\ c(k) \Q\" + shows "I. \P\ EXS k. c(k) \Q\" +using h +by (smt IHoare_def code_exI) + +lemma I_code_extension: + assumes h: "I. \ P \ c \ Q \" + shows "I. \ P \ c ** e \ Q \" + using h + by (smt IHoare_def sep_exec.code_extension) + +lemma I_composition: + assumes h1: "I. \P\ c1 \Q\" + and h2: "I. \Q\ c2 \R\" + shows "I. \P\ c1 ** c2 \R\" + using h1 h2 +by (smt IHoare_def sep_exec.composition) + +lemma pre_condI: + assumes h: "cond \ \p\ c \q\" + shows "\ \* p\ c \q\" +proof(induct rule:HoareI) + case (Pre r s) + hence "cond" "(p \* c \* r) (recse s)" + apply (metis pasrt_def sep_conjD) + by (smt Pre.hyps pasrt_def sep_add_zero sep_conj_commute sep_conj_def) + from h[OF this(1), unfolded Hoare_gen_def, rule_format, OF this(2)] + show ?case . +qed + +lemma I_pre_condI: + assumes h: "cond \ I.\P\ c \Q\" + shows "I.\ \* P\ c \Q\" + using h +by (smt IHoareI condD cond_true_eq2 sep.mult_commute) + +lemma code_condI: + assumes h: "b \ \p\ c \q\" + shows "\p\ **c \q\" +proof(induct rule: HoareI) + case (Pre r s) + hence h1: "b" "(p \* c \* r) (recse s)" + apply (metis condD sep_conjD sep_conj_assoc) + by (smt Pre.hyps condD sep_conj_impl) + from h[OF h1(1), unfolded Hoare_gen_def, rule_format, OF h1(2)] + and h1(1) + show ?case + by (metis (full_types) cond_true_eq1) +qed + +lemma I_code_condI: + assumes h: "b \ I. \P\ c \Q\" + shows "I.\P\ **c \Q\" + using h +by (smt IHoareI condD cond_true_eq2 sep.mult_commute sep_conj_cond1) + +lemma precond_exI: + assumes h:"\x. \p x\ c \q\" + shows "\EXS x. p x\ c \q\" +proof(induct rule:HoareI) + case (Pre r s) + then obtain x where "(p x \* c \* r) (recse s)" + by (unfold pred_ex_def, auto elim!:sep_conjE intro!:sep_conjI) + from h[of x, unfolded Hoare_gen_def, rule_format, OF this] + show ?case . +qed + +lemma I_precond_exI: + assumes h:"\x. I. \P x\ c \Q\" + shows "I.\EXS x. P x\ c \Q\" +proof(induct rule:IHoareI) + case (IPre s' s r cnf) + then obtain x + where "((EXS s.

\* <(s ## s')> \* I (s + s')) \* c \* r) (recse cnf)" + by ( auto elim!:sep_conjE intro!:sep_conjI simp:pred_ex_def pasrt_def sep_conj_ac) + from h[unfolded IHoare_def Hoare_gen_def, rule_format, OF this] + obtain k t + where "(( \* <(t ## s')> \* I (t + s')) \* c \* r) (recse (run (Suc k) cnf))" + by (unfold pred_ex_def, auto elim!:sep_conjE intro!:sep_conjI) + thus ?case + by (auto simp:sep_conj_ac) +qed + +lemma hoare_sep_false: + "\sep_false\ c + \q\" + by(unfold Hoare_gen_def, clarify, simp) + +lemma I_hoare_sep_false: + "I. \sep_false\ c + \Q\" +by (smt IHoareI condD) + +end + +instantiation set :: (type)sep_algebra +begin + +definition set_zero_def: "0 = {}" + +definition plus_set_def: "s1 + s2 = s1 \ s2" + +definition sep_disj_set_def: "sep_disj s1 s2 = (s1 \ s2 = {})" + +lemmas set_ins_def = sep_disj_set_def plus_set_def set_zero_def + +instance + apply(default) + apply(simp add:set_ins_def) + apply(simp add:sep_disj_set_def plus_set_def set_zero_def) + apply (metis inf_commute) + apply(simp add:sep_disj_set_def plus_set_def set_zero_def) + apply(simp add:sep_disj_set_def plus_set_def set_zero_def) + apply (metis sup_commute) + apply(simp add:sep_disj_set_def plus_set_def set_zero_def) + apply (metis (lifting) Un_assoc) + apply(simp add:sep_disj_set_def plus_set_def set_zero_def) + apply (metis (lifting) Int_Un_distrib Un_empty inf_sup_distrib1 sup_eq_bot_iff) + apply(simp add:sep_disj_set_def plus_set_def set_zero_def) + by (metis (lifting) Int_Un_distrib Int_Un_distrib2 sup_eq_bot_iff) +end + +section {* A big operator of infinite separation conjunction *} + +definition "fam_conj I cpt s = (\ p. (s = (\ i \ I. p(i))) \ + (\ i \ I. cpt i (p i)) \ + (\ i \ I. \ j \ I. i \ j \ p(i) ## p(j)))" + +lemma fam_conj_zero_simp: "fam_conj {} cpt = " +proof + fix s + show "fam_conj {} cpt s = () s" + proof + assume "fam_conj {} cpt s" + then obtain p where "s = (\ i \ {}. p(i))" + by (unfold fam_conj_def, auto) + hence "s = {}" by auto + thus "() s" by (metis pasrt_def set_zero_def) + next + assume "() s" + hence eq_s: "s = {}" by (metis pasrt_def set_zero_def) + let ?p = "\ i. {}" + have "(s = (\ i \ {}. ?p(i)))" by (unfold eq_s, auto) + moreover have "(\ i \ {}. cpt i (?p i))" by auto + moreover have "(\ i \ {}. \ j \ {}. i \ j \ ?p(i) ## ?p(j))" by auto + ultimately show "fam_conj {} cpt s" + by (unfold eq_s fam_conj_def, auto) + qed +qed + +lemma fam_conj_disj_simp_pre: + assumes h1: "I = I1 + I2" + and h2: "I1 ## I2" + shows "fam_conj I cpt = (fam_conj I1 cpt \* fam_conj I2 cpt)" +proof + fix s + let ?fm = "fam_conj I cpt" and ?fm1 = "fam_conj I1 cpt" and ?fm2 = "fam_conj I2 cpt" + show "?fm s = (?fm1 \* ?fm2) s" + proof + assume "?fm s" + then obtain p where pre: + "s = (\ i \ I. p(i))" + "(\ i \ I. cpt i (p i))" + "(\ i \ I. \ j \ I. i \ j \ p(i) ## p(j))" + unfolding fam_conj_def by metis + from pre(1) h1 h2 have "s = (\ i \ I1. p(i)) + (\ i \ I2. p(i))" + by (auto simp:set_ins_def) + moreover from pre h1 have "?fm1 (\ i \ I1. p(i))" + by (unfold fam_conj_def, rule_tac x = p in exI, auto simp:set_ins_def) + moreover from pre h1 have "?fm2 (\ i \ I2. p(i))" + by (unfold fam_conj_def, rule_tac x = p in exI, auto simp:set_ins_def) + moreover have "(\ i \ I1. p(i)) ## (\ i \ I2. p(i))" + proof - + { fix x xa xb + assume h: "I1 \ I2 = {}" + "\i\I1 \ I2. \j\I1 \ I2. i \ j \ p i \ p j = {}" + "xa \ I1" "x \ p xa" "xb \ I2" "x \ p xb" + have "p xa \ p xb = {}" + proof(rule h(2)[rule_format]) + from h(1, 3, 5) show "xa \ I1 \ I2" by auto + next + from h(1, 3, 5) show "xb \ I1 \ I2" by auto + next + from h(1, 3, 5) show "xa \ xb" by auto + qed with h(4, 6) have False by auto + } with h1 h2 pre(3) show ?thesis by (auto simp:set_ins_def) + qed + ultimately show "(?fm1 \* ?fm2) s" by (auto intro!:sep_conjI) + next + assume "(?fm1 \* ?fm2) s" + then obtain s1 s2 where pre: + "s = s1 + s2" "s1 ## s2" "?fm1 s1" "?fm2 s2" + by (auto dest!:sep_conjD) + from pre(3) obtain p1 where pre1: + "s1 = (\ i \ I1. p1(i))" + "(\ i \ I1. cpt i (p1 i))" + "(\ i \ I1. \ j \ I1. i \ j \ p1(i) ## p1(j))" + unfolding fam_conj_def by metis + from pre(4) obtain p2 where pre2: + "s2 = (\ i \ I2. p2(i))" + "(\ i \ I2. cpt i (p2 i))" + "(\ i \ I2. \ j \ I2. i \ j \ p2(i) ## p2(j))" + unfolding fam_conj_def by metis + let ?p = "\ i. if i \ I1 then p1 i else p2 i" + from h2 pre(2) + have "s = (\ i \ I. ?p(i))" + apply (unfold h1 pre(1) pre1(1) pre2(1) set_ins_def, auto split:if_splits) + by (metis disjoint_iff_not_equal) + moreover from h2 pre1(2) pre2(2) have "(\ i \ I. cpt i (?p i))" + by (unfold h1 set_ins_def, auto split:if_splits) + moreover from pre1(1, 3) pre2(1, 3) pre(2) h2 + have "(\ i \ I. \ j \ I. i \ j \ ?p(i) ## ?p(j))" + apply (unfold h1 set_ins_def, auto split:if_splits) + by (metis IntI empty_iff) + ultimately show "?fm s" by (unfold fam_conj_def, auto) + qed +qed + +lemma fam_conj_disj_simp: + assumes h: "I1 ## I2" + shows "fam_conj (I1 + I2) cpt = (fam_conj I1 cpt \* fam_conj I2 cpt)" +proof - + from fam_conj_disj_simp_pre[OF _ h, of "I1 + I2"] + show ?thesis by simp +qed + +lemma fam_conj_elm_simp: + assumes h: "i \ I" + shows "fam_conj I cpt = (cpt(i) \* fam_conj (I - {i}) cpt)" +proof + fix s + show "fam_conj I cpt s = (cpt i \* fam_conj (I - {i}) cpt) s" + proof + assume pre: "fam_conj I cpt s" + show "(cpt i \* fam_conj (I - {i}) cpt) s" + proof - + from pre obtain p where pres: + "s = (\ i \ I. p(i))" + "(\ i \ I. cpt i (p i))" + "(\ i \ I. \ j \ I. i \ j \ p(i) ## p(j))" + unfolding fam_conj_def by metis + let ?s = "(\i\(I - {i}). p i)" + from pres(3) h have disj: "(p i) ## ?s" + by (auto simp:set_ins_def) + moreover from pres(1) h have eq_s: "s = (p i) + ?s" + by (auto simp:set_ins_def) + moreover from pres(2) h have "cpt i (p i)" by auto + moreover from pres have "(fam_conj (I - {i}) cpt) ?s" + by (unfold fam_conj_def, rule_tac x = p in exI, auto) + ultimately show ?thesis by (auto intro!:sep_conjI) + qed + next + let ?fam = "fam_conj (I - {i}) cpt" + assume "(cpt i \* ?fam) s" + then obtain s1 s2 where pre: + "s = s1 + s2" "s1 ## s2" "cpt i s1" "?fam s2" + by (auto dest!:sep_conjD) + from pre(4) obtain p where pres: + "s2 = (\ ia \ I - {i}. p(ia))" + "(\ ia \ I - {i}. cpt ia (p ia))" + "(\ ia \ I - {i}. \ j \ I - {i}. ia \ j \ p(ia) ## p(j))" + unfolding fam_conj_def by metis + let ?p = "p(i:=s1)" + from h pres(3) pre(2)[unfolded pres(1)] + have " \ia\I. \j\I. ia \ j \ ?p ia ## ?p j" by (auto simp:set_ins_def) + moreover from pres(1) pre(1) h pre(2) + have "(s = (\ i \ I. ?p(i)))" by (auto simp:set_ins_def split:if_splits) + moreover from pre(3) pres(2) h + have "(\ i \ I. cpt i (?p i))" by (auto simp:set_ins_def split:if_splits) + ultimately show "fam_conj I cpt s" + by (unfold fam_conj_def, auto) + qed +qed + +lemma fam_conj_insert_simp: + assumes h:"i \ I" + shows "fam_conj (insert i I) cpt = (cpt(i) \* fam_conj I cpt)" +proof - + have "i \ insert i I" by auto + from fam_conj_elm_simp[OF this] and h + show ?thesis by auto +qed + +lemmas fam_conj_simps = fam_conj_insert_simp fam_conj_zero_simp + +lemma "fam_conj {0, 2, 3} cpt = (cpt 2 \* cpt (0::int) \* cpt 3)" + by (simp add:fam_conj_simps sep_conj_ac) + +lemma fam_conj_ext_eq: + assumes h: "\ i . i \ I \ f i = g i" + shows "fam_conj I f = fam_conj I g" +proof + fix s + show "fam_conj I f s = fam_conj I g s" + by (unfold fam_conj_def, auto simp:h) +qed + +end + diff -r 77daf1b85cf0 -r a5f5b9336007 thys2/Hoare_tm.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/Hoare_tm.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,5438 @@ +header {* + Separation logic for TM +*} + +theory Hoare_tm +imports Hoare_tm_basis My_block Data_slot +begin + + +section {* Aux lemmas *} + +lemma int_add_C :"x + (y::int) = y + x" + by simp + +lemma int_add_A : "(x + y) + z = x + (y + (z::int))" + by simp + +lemma int_add_LC: "x + (y + (z::int)) = y + (x + z)" + by simp + +lemmas int_add_ac = int_add_A int_add_C int_add_LC + +method_setup prune = {* Scan.succeed (SIMPLE_METHOD' o (K (K prune_params_tac))) *} + "pruning parameters" + +ML {* +structure StepRules = Named_Thms + (val name = @{binding "step"} + val description = "Theorems for hoare rules for every step") +*} + +ML {* +structure FwdRules = Named_Thms + (val name = @{binding "fwd"} + val description = "Theorems for fwd derivation of seperation implication") +*} + +setup {* StepRules.setup *} + +setup {* FwdRules.setup *} + +section {* Operational Semantics of TM *} + +type_synonym tconf = "nat \ (nat \ tm_inst) \ nat \ int \ (int \ Block)" + +fun next_tape :: "taction \ (int \ (int \ Block)) \ (int \ (int \ Block))" +where "next_tape W0 (pos, m) = (pos, m(pos \ Bk))" | + "next_tape W1 (pos, m) = (pos, m(pos \ Oc))" | + "next_tape L (pos, m) = (pos - 1, m)" | + "next_tape R (pos, m) = (pos + 1, m)" + +fun tstep :: "tconf \ tconf" + where "tstep (faults, prog, pc, pos, m) = + (case (prog pc) of + Some ((action0, St pc0), (action1, St pc1)) \ + case (m pos) of + Some Bk \ (faults, prog, pc0, next_tape action0 (pos, m)) | + Some Oc \ (faults, prog, pc1, next_tape action1 (pos, m)) | + None \ (faults + 1, prog, pc, pos, m) + | None \ (faults + 1, prog, pc, pos, m))" + +(* lemma tstep_splits: + "(P (tstep s)) = ((\ faults prog pc pos m action0 pc0 action1 pc1. + s = (faults, prog, pc, pos, m) \ + prog pc = Some ((action0, pc0), (action1, pc1)) \ + m pos = Some Bk \ P (faults, prog, pc0, next_tape action0 (pos, m))) \ + (\ faults prog pc pos m action0 pc0 action1 pc1. + s = (faults, prog, pc, pos, m) \ + prog pc = Some ((action0, pc0), (action1, pc1)) \ + m pos = Some Oc \ P (faults, prog, pc1, next_tape action1 (pos, m))) \ + (\ faults prog pc pos m action0 pc0 action1 pc1. + s = (faults, prog, pc, pos, m) \ + prog pc = Some ((action0, pc0), (action1, pc1)) \ + m pos = None \ P (faults + 1, prog, pc, pos, m)) \ + (\ faults prog pc pos m . + s = (faults, prog, pc, pos, m) \ + prog pc = None \ P (faults + 1, prog, pc, pos, m)) + )" + apply (case_tac s, auto split:option.splits Block.splits) +*) + +definition "tprog_set prog = {TC i inst | i inst. prog i = Some inst}" +definition "tpc_set pc = {TAt pc}" +definition "tmem_set m = {TM i n | i n. m (i) = Some n}" +definition "tpos_set pos = {TPos pos}" +definition "tfaults_set faults = {TFaults faults}" + +lemmas tpn_set_def = tprog_set_def tpc_set_def tmem_set_def tfaults_set_def tpos_set_def + +fun trset_of :: "tconf \ tresource set" + where "trset_of (faults, prog, pc, pos, m) = + tprog_set prog \ tpc_set pc \ tpos_set pos \ tmem_set m \ tfaults_set faults" + +interpretation tm: sep_exec tstep trset_of . + +section {* Hoare logic for TM *} + +abbreviation t_hoare :: + "tassert \ tassert \ tassert \ bool" ("(\(1_)\ / (_)/ \(1_)\)" 50) + where "\ p \ c \ q \ == sep_exec.Hoare_gen tstep trset_of p c q" + +abbreviation it_hoare :: + "(('a::sep_algebra) \ tresource set \ bool) \ + ('a \ bool) \ (tresource set \ bool) \ ('a \ bool) \ bool" + ("(1_).(\(1_)\ / (_)/ \(1_)\)" 50) +where "I. \P\ c \Q\ == sep_exec.IHoare tstep trset_of I P c Q" + +(* +primrec tpg_len :: "tpg \ nat" where + "tpg_len (TInstr ai) = 1" | + "tpg_len (TSeq p1 p2) = tpg_len p1 + tpg_len " | + "tpg_len (TLocal fp) = tpg_len (fp 0)" | + "tpg_len (TLabel l) = 0" *) + +lemma tpg_fold_sg: "tpg_fold [tpg] = tpg" + by (simp add:tpg_fold_def) + +lemma tpg_fold_cons: + assumes h: "tpgs \ []" + shows "tpg_fold (tpg#tpgs) = (tpg; (tpg_fold tpgs))" + using h +proof(induct tpgs arbitrary:tpg) + case (Cons tpg1 tpgs1) + thus ?case + proof(cases "tpgs1 = []") + case True + thus ?thesis by (simp add:tpg_fold_def) + next + case False + show ?thesis + proof - + have eq_1: "butlast (tpg # tpg1 # tpgs1) = tpg # (butlast (tpg1 # tpgs1))" + by simp + from False have eq_2: "last (tpg # tpg1 # tpgs1) = last (tpg1 # tpgs1)" + by simp + have eq_3: "foldr (op ;) (tpg # butlast (tpg1 # tpgs1)) (last (tpg1 # tpgs1)) = + (tpg ; (foldr (op ;) (butlast (tpg1 # tpgs1)) (last (tpg1 # tpgs1))))" + by simp + show ?thesis + apply (subst (1) tpg_fold_def, unfold eq_1 eq_2 eq_3) + by (fold tpg_fold_def, simp) + qed + qed +qed auto + +lemmas tpg_fold_simps = tpg_fold_sg tpg_fold_cons + +lemma tpg_fold_app: + assumes h1: "tpgs1 \ []" + and h2: "tpgs2 \ []" + shows "i:[(tpg_fold (tpgs1 @ tpgs2))]:j = i:[(tpg_fold (tpgs1); tpg_fold tpgs2)]:j" + using h1 h2 +proof(induct tpgs1 arbitrary: i j tpgs2) + case (Cons tpg1' tpgs1') + thus ?case (is "?L = ?R") + proof(cases "tpgs1' = []") + case False + from h2 have "(tpgs1' @ tpgs2) \ []" + by (metis Cons.prems(2) Nil_is_append_conv) + have "?L = (i :[ tpg_fold (tpg1' # (tpgs1' @ tpgs2)) ]: j )" by simp + also have "\ = (i:[(tpg1'; (tpg_fold (tpgs1' @ tpgs2)))]:j )" + by (simp add:tpg_fold_cons[OF `(tpgs1' @ tpgs2) \ []`]) + also have "\ = ?R" + proof - + have "(EXS j'. i :[ tpg1' ]: j' \* j' :[ tpg_fold (tpgs1' @ tpgs2) ]: j) = + (EXS j'. (EXS j'a. i :[ tpg1' ]: j'a \* j'a :[ tpg_fold tpgs1' ]: j') \* + j' :[ tpg_fold tpgs2 ]: j)" + proof(default+) + fix s + assume "(EXS j'. i :[ tpg1' ]: j' \* j' :[ tpg_fold (tpgs1' @ tpgs2) ]: j) s" + thus "(EXS j'. (EXS j'a. i :[ tpg1' ]: j'a \* j'a :[ tpg_fold tpgs1' ]: j') \* + j' :[ tpg_fold tpgs2 ]: j) s" + proof(elim EXS_elim) + fix j' + assume "(i :[ tpg1' ]: j' \* j' :[ tpg_fold (tpgs1' @ tpgs2) ]: j) s" + from this[unfolded Cons(1)[OF False Cons(3)] tassemble_to.simps] + show "(EXS j'. (EXS j'a. i :[ tpg1' ]: j'a \* j'a :[ tpg_fold tpgs1' ]: j') \* + j' :[ tpg_fold tpgs2 ]: j) s" + proof(elim sep_conjE EXS_elim) + fix x y j'a xa ya + assume h: "(i :[ tpg1' ]: j') x" + "x ## y" "s = x + y" + "(j' :[ tpg_fold tpgs1' ]: j'a) xa" + "(j'a :[ tpg_fold tpgs2 ]: j) ya" + " xa ## ya" "y = xa + ya" + show "(EXS j'. (EXS j'a. i :[ tpg1' ]: j'a \* + j'a :[ tpg_fold tpgs1' ]: j') \* j' :[ tpg_fold tpgs2 ]: j) s" + (is "(EXS j'. (?P j' \* ?Q j')) s") + proof(rule EXS_intro[where x = "j'a"]) + from `(j'a :[ tpg_fold tpgs2 ]: j) ya` have "(?Q j'a) ya" . + moreover have "(?P j'a) (x + xa)" + proof(rule EXS_intro[where x = j']) + have "x + xa = x + xa" by simp + moreover from `x ## y` `y = xa + ya` `xa ## ya` + have "x ## xa" by (metis sep_disj_addD) + moreover note `(i :[ tpg1' ]: j') x` `(j' :[ tpg_fold tpgs1' ]: j'a) xa` + ultimately show "(i :[ tpg1' ]: j' \* j' :[ tpg_fold tpgs1' ]: j'a) (x + xa)" + by (auto intro!:sep_conjI) + qed + moreover from `x##y` `y = xa + ya` `xa ## ya` + have "(x + xa) ## ya" + by (metis sep_disj_addI1 sep_disj_commuteI) + moreover from `s = x + y` `y = xa + ya` + have "s = (x + xa) + ya" + by (metis h(2) h(6) sep_add_assoc sep_disj_addD1 sep_disj_addD2) + ultimately show "(?P j'a \* ?Q j'a) s" + by (auto intro!:sep_conjI) + qed + qed + qed + next + fix s + assume "(EXS j'. (EXS j'a. i :[ tpg1' ]: j'a \* j'a :[ tpg_fold tpgs1' ]: j') \* + j' :[ tpg_fold tpgs2 ]: j) s" + thus "(EXS j'. i :[ tpg1' ]: j' \* j' :[ tpg_fold (tpgs1' @ tpgs2) ]: j) s" + proof(elim EXS_elim sep_conjE) + fix j' x y j'a xa ya + assume h: "(j' :[ tpg_fold tpgs2 ]: j) y" + "x ## y" "s = x + y" "(i :[ tpg1' ]: j'a) xa" + "(j'a :[ tpg_fold tpgs1' ]: j') ya" "xa ## ya" "x = xa + ya" + show "(EXS j'. i :[ tpg1' ]: j' \* j' :[ tpg_fold (tpgs1' @ tpgs2) ]: j) s" + proof(rule EXS_intro[where x = j'a]) + from `(i :[ tpg1' ]: j'a) xa` have "(i :[ tpg1' ]: j'a) xa" . + moreover have "(j'a :[ tpg_fold (tpgs1' @ tpgs2) ]: j) (ya + y)" + proof(unfold Cons(1)[OF False Cons(3)] tassemble_to.simps) + show "(EXS j'. j'a :[ tpg_fold tpgs1' ]: j' \* j' :[ tpg_fold tpgs2 ]: j) (ya + y)" + proof(rule EXS_intro[where x = "j'"]) + from `x ## y` `x = xa + ya` `xa ## ya` + have "ya ## y" by (metis sep_add_disjD) + moreover have "ya + y = ya + y" by simp + moreover note `(j'a :[ tpg_fold tpgs1' ]: j') ya` + `(j' :[ tpg_fold tpgs2 ]: j) y` + ultimately show "(j'a :[ tpg_fold tpgs1' ]: j' \* + j' :[ tpg_fold tpgs2 ]: j) (ya + y)" + by (auto intro!:sep_conjI) + qed + qed + moreover from `s = x + y` `x = xa + ya` + have "s = xa + (ya + y)" + by (metis h(2) h(6) sep_add_assoc sep_add_disjD) + moreover from `xa ## ya` `x ## y` `x = xa + ya` + have "xa ## (ya + y)" by (metis sep_disj_addI3) + ultimately show "(i :[ tpg1' ]: j'a \* j'a :[ tpg_fold (tpgs1' @ tpgs2) ]: j) s" + by (auto intro!:sep_conjI) + qed + qed + qed + thus ?thesis + by (simp add:tassemble_to.simps, unfold tpg_fold_cons[OF False], + unfold tassemble_to.simps, simp) + qed + finally show ?thesis . + next + case True + thus ?thesis + by (simp add:tpg_fold_cons[OF Cons(3)] tpg_fold_sg) + qed +qed auto + + +subsection {* Assertions and program logic for this assembly language *} + +definition "st l = sg (tpc_set l)" +definition "ps p = sg (tpos_set p)" +definition "tm a v =sg ({TM a v})" +definition "one i = tm i Oc" +definition "zero i= tm i Bk" +definition "any i = (EXS v. tm i v)" + +declare trset_of.simps[simp del] + +lemma stimes_sgD: "(sg x ** q) s \ q (s - x) \ x \ s" + apply(erule_tac sep_conjE) + apply(unfold set_ins_def sg_def) + by (metis Diff_Int2 Diff_Int_distrib2 Diff_Un Diff_cancel + Diff_empty Diff_idemp Diff_triv Int_Diff Un_Diff + Un_Diff_cancel inf_commute inf_idem sup_bot_right sup_commute sup_ge2) + +lemma stD: "(st i ** r) (trset_of (ft, prog, i', pos, mem)) + \ i' = i" +proof - + assume "(st i ** r) (trset_of (ft, prog, i', pos, mem))" + from stimes_sgD [OF this[unfolded st_def], unfolded trset_of.simps] + have "tpc_set i \ tprog_set prog \ tpc_set i' \ tpos_set pos \ + tmem_set mem \ tfaults_set ft" by auto + thus ?thesis + by (unfold tpn_set_def, auto) +qed + +lemma psD: "(ps p ** r) (trset_of (ft, prog, i', pos, mem)) + \ pos = p" +proof - + assume "(ps p ** r) (trset_of (ft, prog, i', pos, mem))" + from stimes_sgD [OF this[unfolded ps_def], unfolded trset_of.simps] + have "tpos_set p \ tprog_set prog \ tpc_set i' \ + tpos_set pos \ tmem_set mem \ tfaults_set ft" + by simp + thus ?thesis + by (unfold tpn_set_def, auto) +qed + +lemma codeD: "(st i ** sg {TC i inst} ** r) (trset_of (ft, prog, i', pos, mem)) + \ prog i = Some inst" +proof - + assume "(st i ** sg {TC i inst} ** r) (trset_of (ft, prog, i', pos, mem))" + thus ?thesis + apply(unfold sep_conj_def set_ins_def sg_def trset_of.simps tpn_set_def) + by auto +qed + +lemma memD: "((tm a v) ** r) (trset_of (ft, prog, i, pos, mem)) \ mem a = Some v" +proof - + assume "((tm a v) ** r) (trset_of (ft, prog, i, pos, mem))" + from stimes_sgD[OF this[unfolded trset_of.simps tpn_set_def tm_def]] + have "{TM a v} \ {TC i inst |i inst. prog i = Some inst} \ {TAt i} \ + {TPos pos} \ {TM i n |i n. mem i = Some n} \ {TFaults ft}" by simp + thus ?thesis by auto +qed + +lemma t_hoare_seq: + "\\ i j. \st i ** p\ i:[c1]:j \st j ** q\; + \ j k. \st j ** q\ j:[c2]:k \st k ** r\\ \ \st i ** p\ i:[(c1 ; c2)]:k \st k ** r\" +proof - + assume h: "\ i j. \st i ** p\ i:[c1]:j \st j ** q\" + "\ j k. \st j ** q\ j:[c2]:k \st k ** r\" + show "\st i ** p\ i:[(c1 ; c2)]:k \st k ** r\" + proof(subst tassemble_to.simps, rule tm.code_exI) + fix j' + show "\st i \* p\ i :[ c1 ]: j' \* j' :[ c2 ]: k \st k \* r\" + proof(rule tm.composition) + from h(1) show "\st i \* p\ i :[ c1 ]: j' \ st j' \* q \" by auto + next + from h(2) show "\st j' \* q\ j' :[ c2 ]: k \st k \* r\" by auto + qed + qed +qed + +lemma t_hoare_seq1: + "\\j'. \st i ** p\ i:[c1]:j' \st j' ** q\; + \j'. \st j' ** q\ j':[c2]:k \st k' ** r\\ \ + \st i ** p\ i:[(c1 ; c2)]:k \st k' ** r\" +proof - + assume h: "\j'. \st i \* p\ i :[ c1 ]: j' \st j' \* q\" + "\j'. \st j' \* q\ j' :[ c2 ]: k \st k' \* r\" + show "\st i \* p\ i :[ (c1 ; c2) ]: k \st k' \* r\" + proof(subst tassemble_to.simps, rule tm.code_exI) + fix j' + show "\st i \* p\ i :[ c1 ]: j' \* j' :[ c2 ]: k \st k' \* r\" + proof(rule tm.composition) + from h(1) show "\st i \* p\ i :[ c1 ]: j' \ st j' \* q \" by auto + next + from h(2) show " \st j' \* q\ j' :[ c2 ]: k \st k' \* r\" by auto + qed + qed +qed + +lemma t_hoare_seq2: + assumes h: "\j. \st i ** p\ i:[c1]:j \st k' \* r\" + shows "\st i ** p\ i:[(c1 ; c2)]:j \st k' ** r\" + apply (unfold tassemble_to.simps, rule tm.code_exI) + by (rule tm.code_extension, rule h) + +lemma t_hoare_local: + assumes h: "(\ (l::nat). \st i \* p\ i :[ c l ]: j \st k \* q\)" + shows "\st i ** p\ i:[TLocal c]:j \st k ** q\" using h + by (unfold tassemble_to.simps, intro tm.code_exI, case_tac l, simp) + +lemma t_hoare_label: + "(\l. (l::nat) = i \ \st l \* p\ l :[ c (l::tstate) ]: j \st k \* q\) \ + \st i \* p \ + i:[(TLabel l; c l)]:j + \st k \* q\" + by (cases l, unfold tassemble_to.simps, intro tm.code_exI tm.code_condI, clarify, auto) + +primrec t_split_cmd :: "tpg \ (tpg \ tpg option)" + where "t_split_cmd (\inst) = (\inst, None)" | + "t_split_cmd (TLabel l) = (TLabel l, None)" | + "t_split_cmd (TSeq c1 c2) = (case (t_split_cmd c2) of + (c21, Some c22) \ (TSeq c1 c21, Some c22) | + (c21, None) \ (c1, Some c21))" | + "t_split_cmd (TLocal c) = (TLocal c, None)" + +definition "t_last_cmd tpg = (snd (t_split_cmd tpg))" + +declare t_last_cmd_def [simp] + +definition "t_blast_cmd tpg = (fst (t_split_cmd tpg))" + +declare t_blast_cmd_def [simp] + +lemma "t_last_cmd (c1; c2; (TLabel l)) = (Some (TLabel l))" + by simp + +lemma "t_blast_cmd (c1; c2; TLabel l) = (c1; c2)" + by simp + +lemma t_split_cmd_eq: + assumes "t_split_cmd c = (c1, Some c2)" + shows "(i:[c]:j) = (i:[(c1; c2)]:j)" + using assms +proof(induct c arbitrary: c1 c2 i j) + case (TSeq cx cy) + from `t_split_cmd (cx ; cy) = (c1, Some c2)` + show ?case + apply (simp split:prod.splits option.splits) + apply (cases cy, auto split:prod.splits option.splits) + proof - + fix a + assume h: "t_split_cmd cy = (a, Some c2)" + show "i :[ (cx ; cy) ]: j = i :[ ((cx ; a) ; c2) ]: j" + apply (simp only: tassemble_to.simps(2) TSeq(2)[OF h] sep_conj_exists) + apply (simp add:sep_conj_ac) + by (simp add:pred_ex_def, default, auto) + qed +qed auto + +lemma t_hoare_label_last_pre: + fixes l + assumes h1: "t_split_cmd c = (c', Some (TLabel l))" + and h2: "l = j \ \p\ i:[c']:j \q\" + shows "\p\ i:[c]:j \q\" +by (cases l, unfold t_split_cmd_eq[OF h1], + simp only:tassemble_to.simps sep_conj_cond, + intro tm.code_exI tm.code_condI, insert h2, auto) + +lemma t_hoare_label_last: + fixes l + assumes h1: "t_last_cmd c = Some (TLabel l)" + and h2: "l = j \ \p\ i:[t_blast_cmd c]:j \q\" + shows "\p\ i:[c]:j \q\" +proof - + have "t_split_cmd c = (t_blast_cmd c, t_last_cmd c)" + by simp + from t_hoare_label_last_pre[OF this[unfolded h1]] h2 + show ?thesis by auto +qed + +subsection {* Basic assertions for TM *} + +function ones :: "int \ int \ tassert" where + "ones i j = (if j < i then <(i = j + 1)> else + (one i) ** ones (i + 1) j)" +by auto +termination by (relation "measure(\ (i::int, j). nat (j - i + 1))") auto + +lemma ones_base_simp: "j < i \ ones i j = <(i = j + 1)>" + by simp + +lemma ones_step_simp: "\ j < i \ ones i j = ((one i) ** ones (i + 1) j)" + by simp + +lemmas ones_simps = ones_base_simp ones_step_simp + +declare ones.simps [simp del] ones_simps [simp] + +lemma ones_induct [case_names Base Step]: + fixes P i j + assumes h1: "\ i j. j < i \ P i j (<(i = j + (1::int))>)" + and h2: "\ i j. \\ j < i; P (i + 1) j (ones (i + 1) j)\ \ P i j ((one i) ** ones (i + 1) j)" + shows "P i j (ones i j)" +proof(induct i j rule:ones.induct) + fix i j + assume ih: "(\ j < i \ P (i + 1) j (ones (i + 1) j))" + show "P i j (ones i j)" + proof(cases "j < i") + case True + with h1 [OF True] + show ?thesis by auto + next + case False + from h2 [OF False] and ih[OF False] + have "P i j (one i \* ones (i + 1) j)" by blast + with False show ?thesis by auto + qed +qed + +function ones' :: "int \ int \ tassert" where + "ones' i j = (if j < i then <(i = j + 1)> else + ones' i (j - 1) ** one j)" +by auto +termination by (relation "measure(\ (i::int, j). nat (j - i + 1))") auto + +lemma ones_rev: "\ j < i \ (ones i j) = ((ones i (j - 1)) ** one j)" +proof(induct i j rule:ones_induct) + case Base + thus ?case by auto +next + case (Step i j) + show ?case + proof(cases "j < i + 1") + case True + with Step show ?thesis + by simp + next + case False + with Step show ?thesis + by (auto simp:sep_conj_ac) + qed +qed + +lemma ones_rev_induct [case_names Base Step]: + fixes P i j + assumes h1: "\ i j. j < i \ P i j (<(i = j + (1::int))>)" + and h2: "\ i j. \\ j < i; P i (j - 1) (ones i (j - 1))\ \ P i j ((ones i (j - 1)) ** one j)" + shows "P i j (ones i j)" +proof(induct i j rule:ones'.induct) + fix i j + assume ih: "(\ j < i \ P i (j - 1) (ones i (j - 1)))" + show "P i j (ones i j)" + proof(cases "j < i") + case True + with h1 [OF True] + show ?thesis by auto + next + case False + from h2 [OF False] and ih[OF False] + have "P i j (ones i (j - 1) \* one j)" by blast + with ones_rev and False + show ?thesis + by simp + qed +qed + +function zeros :: "int \ int \ tassert" where + "zeros i j = (if j < i then <(i = j + 1)> else + (zero i) ** zeros (i + 1) j)" +by auto +termination by (relation "measure(\ (i::int, j). nat (j - i + 1))") auto + +lemma zeros_base_simp: "j < i \ zeros i j = <(i = j + 1)>" + by simp + +lemma zeros_step_simp: "\ j < i \ zeros i j = ((zero i) ** zeros (i + 1) j)" + by simp + +declare zeros.simps [simp del] +lemmas zeros_simps [simp] = zeros_base_simp zeros_step_simp + +lemma zeros_induct [case_names Base Step]: + fixes P i j + assumes h1: "\ i j. j < i \ P i j (<(i = j + (1::int))>)" + and h2: "\ i j. \\ j < i; P (i + 1) j (zeros (i + 1) j)\ \ + P i j ((zero i) ** zeros (i + 1) j)" + shows "P i j (zeros i j)" +proof(induct i j rule:zeros.induct) + fix i j + assume ih: "(\ j < i \ P (i + 1) j (zeros (i + 1) j))" + show "P i j (zeros i j)" + proof(cases "j < i") + case True + with h1 [OF True] + show ?thesis by auto + next + case False + from h2 [OF False] and ih[OF False] + have "P i j (zero i \* zeros (i + 1) j)" by blast + with False show ?thesis by auto + qed +qed + +lemma zeros_rev: "\ j < i \ (zeros i j) = ((zeros i (j - 1)) ** zero j)" +proof(induct i j rule:zeros_induct) + case (Base i j) + thus ?case by auto +next + case (Step i j) + show ?case + proof(cases "j < i + 1") + case True + with Step show ?thesis by auto + next + case False + with Step show ?thesis by (auto simp:sep_conj_ac) + qed +qed + +lemma zeros_rev_induct [case_names Base Step]: + fixes P i j + assumes h1: "\ i j. j < i \ P i j (<(i = j + (1::int))>)" + and h2: "\ i j. \\ j < i; P i (j - 1) (zeros i (j - 1))\ \ + P i j ((zeros i (j - 1)) ** zero j)" + shows "P i j (zeros i j)" +proof(induct i j rule:ones'.induct) + fix i j + assume ih: "(\ j < i \ P i (j - 1) (zeros i (j - 1)))" + show "P i j (zeros i j)" + proof(cases "j < i") + case True + with h1 [OF True] + show ?thesis by auto + next + case False + from h2 [OF False] and ih[OF False] + have "P i j (zeros i (j - 1) \* zero j)" by blast + with zeros_rev and False + show ?thesis by simp + qed +qed + +definition "rep i j k = ((ones i (i + (int k))) ** <(j = i + int k)>)" + +fun reps :: "int \ int \ nat list\ tassert" + where + "reps i j [] = <(i = j + 1)>" | + "reps i j [k] = (ones i (i + int k) ** <(j = i + int k)>)" | + "reps i j (k # ks) = (ones i (i + int k) ** zero (i + int k + 1) ** reps (i + int k + 2) j ks)" + +lemma reps_simp3: "ks \ [] \ + reps i j (k # ks) = (ones i (i + int k) ** zero (i + int k + 1) ** reps (i + int k + 2) j ks)" + by (cases ks, auto) + +definition "reps_sep_len ks = (if length ks \ 1 then 0 else (length ks) - 1)" + +definition "reps_ctnt_len ks = (\ k \ ks. (1 + k))" + +definition "reps_len ks = (reps_sep_len ks) + (reps_ctnt_len ks)" + +definition "splited xs ys zs = ((xs = ys @ zs) \ (ys \ []) \ (zs \ []))" + +lemma list_split: + assumes h: "k # ks = ys @ zs" + and h1: "ys \ []" + shows "(ys = [k] \ zs = ks) \ (\ ys'. ys' \ [] \ ys = k # ys' \ ks = ys' @ zs)" +proof(cases ys) + case Nil + with h1 show ?thesis by auto +next + case (Cons y' ys') + with h have "k#ks = y'#(ys' @ zs)" by simp + hence hh: "y' = k" "ks = ys' @ zs" by auto + show ?thesis + proof(cases "ys' = []") + case False + show ?thesis + proof(rule disjI2) + show " \ys'. ys' \ [] \ ys = k # ys' \ ks = ys' @ zs" + proof(rule exI[where x = ys']) + from False hh Cons show "ys' \ [] \ ys = k # ys' \ ks = ys' @ zs" by auto + qed + qed + next + case True + show ?thesis + proof(rule disjI1) + from hh True Cons + show "ys = [k] \ zs = ks" by auto + qed + qed +qed + +lemma splited_cons[elim_format]: + assumes h: "splited (k # ks) ys zs" + shows "(ys = [k] \ zs = ks) \ (\ ys'. ys' \ [] \ ys = k # ys' \ splited ks ys' zs)" +proof - + from h have "k # ks = ys @ zs" "ys \ [] " unfolding splited_def by auto + from list_split[OF this] + have "ys = [k] \ zs = ks \ (\ys'. ys' \ [] \ ys = k # ys' \ ks = ys' @ zs)" . + thus ?thesis + proof + assume " ys = [k] \ zs = ks" thus ?thesis by auto + next + assume "\ys'. ys' \ [] \ ys = k # ys' \ ks = ys' @ zs" + then obtain ys' where hh: "ys' \ []" "ys = k # ys'" "ks = ys' @ zs" by auto + show ?thesis + proof(rule disjI2) + show "\ys'. ys' \ [] \ ys = k # ys' \ splited ks ys' zs" + proof(rule exI[where x = ys']) + from h have "zs \ []" by (unfold splited_def, simp) + with hh(1) hh(3) + have "splited ks ys' zs" + by (unfold splited_def, auto) + with hh(1) hh(2) show "ys' \ [] \ ys = k # ys' \ splited ks ys' zs" by auto + qed + qed + qed +qed + +lemma splited_cons_elim: + assumes h: "splited (k # ks) ys zs" + and h1: "\ys = [k]; zs = ks\ \ P" + and h2: "\ ys'. \ys' \ []; ys = k#ys'; splited ks ys' zs\ \ P" + shows P +proof(rule splited_cons[OF h]) + assume "ys = [k] \ zs = ks \ (\ys'. ys' \ [] \ ys = k # ys' \ splited ks ys' zs)" + thus P + proof + assume hh: "ys = [k] \ zs = ks" + show P + proof(rule h1) + from hh show "ys = [k]" by simp + next + from hh show "zs = ks" by simp + qed + next + assume "\ys'. ys' \ [] \ ys = k # ys' \ splited ks ys' zs" + then obtain ys' where hh: "ys' \ []" "ys = k # ys'" "splited ks ys' zs" by auto + from h2[OF this] + show P . + qed +qed + +lemma list_nth_expand: + "i < length xs \ xs = take i xs @ [xs!i] @ drop (Suc i) xs" + by (metis Cons_eq_appendI append_take_drop_id drop_Suc_conv_tl eq_Nil_appendI) + +lemma reps_len_nil: "reps_len [] = 0" + by (auto simp:reps_len_def reps_sep_len_def reps_ctnt_len_def) + +lemma reps_len_sg: "reps_len [k] = 1 + k" + by (auto simp:reps_len_def reps_sep_len_def reps_ctnt_len_def) + +lemma reps_len_cons: "ks \ [] \ (reps_len (k # ks)) = 2 + k + reps_len ks " +proof(induct ks arbitrary:k) + case (Cons n ns) + thus ?case + by(cases "ns = []", + auto simp:reps_len_def reps_sep_len_def reps_ctnt_len_def) +qed auto + +lemma reps_len_correct: + assumes h: "(reps i j ks) s" + shows "j = i + int (reps_len ks) - 1" + using h +proof(induct ks arbitrary:i j s) + case Nil + thus ?case + by (auto simp:reps_len_nil pasrt_def) +next + case (Cons n ns) + thus ?case + proof(cases "ns = []") + case False + from Cons(2)[unfolded reps_simp3[OF False]] + obtain s' where "(reps (i + int n + 2) j ns) s'" + by (auto elim!:sep_conjE) + from Cons.hyps[OF this] + show ?thesis + by (unfold reps_len_cons[OF False], simp) + next + case True + with Cons(2) + show ?thesis + by (auto simp:reps_len_sg pasrt_def) + qed +qed + +lemma reps_len_expand: + shows "(EXS j. (reps i j ks)) = (reps i (i + int (reps_len ks) - 1) ks)" +proof(default+) + fix s + assume "(EXS j. reps i j ks) s" + with reps_len_correct show "reps i (i + int (reps_len ks) - 1) ks s" + by (auto simp:pred_ex_def) +next + fix s + assume "reps i (i + int (reps_len ks) - 1) ks s" + thus "(EXS j. reps i j ks) s" by (auto simp:pred_ex_def) +qed + +lemma reps_len_expand1: "(EXS j. (reps i j ks \* r)) = (reps i (i + int (reps_len ks) - 1) ks \* r)" +by (metis reps_len_def reps_len_expand sep.mult_commute sep_conj_exists1) + +lemma reps_splited: + assumes h: "splited xs ys zs" + shows "reps i j xs = (reps i (i + int (reps_len ys) - 1) ys \* + zero (i + int (reps_len ys)) \* + reps (i + int (reps_len ys) + 1) j zs)" + using h +proof(induct xs arbitrary: i j ys zs) + case Nil + thus ?case + by (unfold splited_def, auto) +next + case (Cons k ks) + from Cons(2) + show ?case + proof(rule splited_cons_elim) + assume h: "ys = [k]" "zs = ks" + with Cons(2) + show ?thesis + proof(cases "ks = []") + case True + with Cons(2) have False + by (simp add:splited_def, cases ys, auto) + thus ?thesis by auto + next + case False + have ss: "i + int k + 1 = i + (1 + int k)" + "i + int k + 2 = 2 + (i + int k)" by auto + show ?thesis + by (unfold h(1), unfold reps_simp3[OF False] reps.simps(2) + reps_len_sg, auto simp:sep_conj_ac, + unfold ss h(2), simp) + qed + next + fix ys' + assume h: "ys' \ []" "ys = k # ys'" "splited ks ys' zs" + hence nnks: "ks \ []" + by (unfold splited_def, auto) + have ss: "i + int k + 2 + int (reps_len ys') = + i + (2 + (int k + int (reps_len ys')))" by auto + show ?thesis + by (simp add: reps_simp3[OF nnks] reps_simp3[OF h(1)] + Cons(1)[OF h(3)] h(2) + reps_len_cons[OF h(1)] + sep_conj_ac, subst ss, simp) + qed +qed + +subsection {* A theory of list extension *} + +definition "list_ext n xs = xs @ replicate ((Suc n) - length xs) 0" + +lemma list_ext_len_eq: "length (list_ext a xs) = length xs + (Suc a - length xs)" + by (metis length_append length_replicate list_ext_def) + +lemma list_ext_len: "a < length (list_ext a xs)" + by (unfold list_ext_len_eq, auto) + +lemma list_ext_lt: "a < length xs \ list_ext a xs = xs" + by (smt append_Nil2 list_ext_def replicate_0) + +lemma list_ext_lt_get: + assumes h: "a' < length xs" + shows "(list_ext a xs)!a' = xs!a'" +proof(cases "a < length xs") + case True + with h + show ?thesis by (auto simp:list_ext_lt) +next + case False + with h show ?thesis + apply (unfold list_ext_def) + by (metis nth_append) +qed + +lemma list_ext_get_upd: "((list_ext a xs)[a:=v])!a = v" + by (metis list_ext_len nth_list_update_eq) + +lemma nth_app: "length xs \ a \ (xs @ ys)!a = ys!(a - length xs)" + by (metis not_less nth_append) + +lemma pred_exI: + assumes "(P(x) \* r) s" + shows "((EXS x. P(x)) \* r) s" + by (metis assms pred_ex_def sep_globalise) + +lemma list_ext_tail: + assumes h1: "length xs \ a" + and h2: "length xs \ a'" + and h3: "a' \ a" + shows "(list_ext a xs)!a' = 0" +proof - + from h1 h2 + have "a' - length xs < length (replicate (Suc a - length xs) 0)" + by (metis diff_less_mono h3 le_imp_less_Suc length_replicate) + moreover from h1 have "replicate (Suc a - length xs) 0 \ []" + by (smt empty_replicate) + ultimately have "(replicate (Suc a - length xs) 0)!(a' - length xs) = 0" + by (metis length_replicate nth_replicate) + moreover have "(xs @ (replicate (Suc a - length xs) 0))!a' = + (replicate (Suc a - length xs) 0)!(a' - length xs)" + by (rule nth_app[OF h2]) + ultimately show ?thesis + by (auto simp:list_ext_def) +qed + +lemmas list_ext_simps = list_ext_lt_get list_ext_lt list_ext_len list_ext_len_eq + +lemma listsum_upd_suc: + "a < length ks \ listsum (map Suc (ks[a := Suc (ks ! a)]))= (Suc (listsum (map Suc ks)))" +by (smt Ex_list_of_length Skolem_list_nth elem_le_listsum_nat + length_induct length_list_update length_map length_splice + list_eq_iff_nth_eq list_ext_get_upd list_ext_lt_get list_update_beyond + list_update_id list_update_overwrite list_update_same_conv list_update_swap + listsum_update_nat map_eq_imp_length_eq map_update neq_if_length_neq + nth_equalityI nth_list_update nth_list_update_eq nth_list_update_neq nth_map reps_sep_len_def) + +lemma reps_len_suc: + assumes "a < length ks" + shows "reps_len (ks[a:=Suc(ks!a)]) = 1 + reps_len ks" +proof(cases "length ks \ 1") + case True + with `a < length ks` + obtain k where "ks = [k]" "a = 0" + by (smt length_0_conv length_Suc_conv) + thus ?thesis + apply (unfold `ks = [k]` `a = 0`) + by (unfold reps_len_def reps_sep_len_def reps_ctnt_len_def, auto) +next + case False + have "Suc = (op + (Suc 0))" + by (default, auto) + with listsum_upd_suc[OF `a < length ks`] False + show ?thesis + by (unfold reps_len_def reps_sep_len_def reps_ctnt_len_def, simp) +qed + +lemma ks_suc_len: + assumes h1: "(reps i j ks) s" + and h2: "ks!a = v" + and h3: "a < length ks" + and h4: "(reps i j' (ks[a:=Suc v])) s'" + shows "j' = j + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1" +proof - + from reps_len_correct[OF h1, unfolded list_ext_len_eq] + reps_len_correct[OF h4, unfolded list_ext_len_eq] + reps_len_suc[OF `a < length ks`] h2 h3 + show ?thesis + by (unfold list_ext_lt[OF `a < length ks`], auto) +qed + +lemma ks_ext_len: + assumes h1: "(reps i j ks) s" + and h3: "length ks \ a" + and h4: "reps i j' (list_ext a ks) s'" + shows "j' = j + int (reps_len (list_ext a ks)) - int (reps_len ks)" +proof - + from reps_len_correct[OF h1, unfolded list_ext_len_eq] + and reps_len_correct[OF h4, unfolded list_ext_len_eq] + h3 + show ?thesis by auto +qed + +definition + "reps' i j ks = + (if ks = [] then (<(i = j + 1)>) else (reps i (j - 1) ks \* zero j))" + +lemma reps'_expand: + assumes h: "ks \ []" + shows "(EXS j. reps' i j ks) = reps' i (i + int (reps_len ks)) ks" +proof - + show ?thesis + proof(default+) + fix s + assume "(EXS j. reps' i j ks) s" + with h have "(EXS j. reps i (j - 1) ks \* zero j) s" + by (simp add:reps'_def) + hence "(reps i (i + int (reps_len ks) - 1) ks \* zero (i + int (reps_len ks))) s" + proof(elim EXS_elim) + fix j + assume "(reps i (j - 1) ks \* zero j) s" + then obtain s1 s2 where "s = s1 + s2" "s1 ## s2" "reps i (j - 1) ks s1" "zero j s2" + by (auto elim!:sep_conjE) + from reps_len_correct[OF this(3)] + have "j = i + int (reps_len ks)" by auto + with `reps i (j - 1) ks s1` have "reps i (i + int (reps_len ks) - 1) ks s1" + by simp + moreover from `j = i + int (reps_len ks)` and `zero j s2` + have "zero (i + int (reps_len ks)) s2" by auto + ultimately show "(reps i (i + int (reps_len ks) - 1) ks \* zero (i + int (reps_len ks))) s" + using `s = s1 + s2` `s1 ## s2` by (auto intro!:sep_conjI) + qed + thus "reps' i (i + int (reps_len ks)) ks s" + by (simp add:h reps'_def) + next + fix s + assume "reps' i (i + int (reps_len ks)) ks s" + thus "(EXS j. reps' i j ks) s" + by (auto intro!:EXS_intro) + qed +qed + +lemma reps'_len_correct: + assumes h: "(reps' i j ks) s" + and h1: "ks \ []" + shows "(j = i + int (reps_len ks))" +proof - + from h1 have "reps' i j ks s = (reps i (j - 1) ks \* zero j) s" by (simp add:reps'_def) + from h[unfolded this] + obtain s' where "reps i (j - 1) ks s'" + by (auto elim!:sep_conjE) + from reps_len_correct[OF this] + show ?thesis by simp +qed + +lemma reps'_append: + "reps' i j (ks1 @ ks2) = (EXS m. (reps' i (m - 1) ks1 \* reps' m j ks2))" +proof(cases "ks1 = []") + case True + thus ?thesis + by (auto simp:reps'_def pred_ex_def pasrt_def set_ins_def sep_conj_def) +next + case False + show ?thesis + proof(cases "ks2 = []") + case False + from `ks1 \ []` and `ks2 \ []` + have "splited (ks1 @ ks2) ks1 ks2" by (auto simp:splited_def) + from reps_splited[OF this, of i "j - 1"] + have eq_1: "reps i (j - 1) (ks1 @ ks2) = + (reps i (i + int (reps_len ks1) - 1) ks1 \* + zero (i + int (reps_len ks1)) \* + reps (i + int (reps_len ks1) + 1) (j - 1) ks2)" . + from `ks1 \ []` + have eq_2: "reps' i j (ks1 @ ks2) = (reps i (j - 1) (ks1 @ ks2) \* zero j)" + by (unfold reps'_def, simp) + show ?thesis + proof(default+) + fix s + assume "reps' i j (ks1 @ ks2) s" + show "(EXS m. reps' i (m - 1) ks1 \* reps' m j ks2) s" + proof(rule EXS_intro[where x = "i + int(reps_len ks1) + 1"]) + from `reps' i j (ks1 @ ks2) s`[unfolded eq_2 eq_1] + and `ks1 \ []` `ks2 \ []` + show "(reps' i (i + int (reps_len ks1) + 1 - 1) ks1 \* + reps' (i + int (reps_len ks1) + 1) j ks2) s" + by (unfold reps'_def, simp, sep_cancel+) + qed + next + fix s + assume "(EXS m. reps' i (m - 1) ks1 \* reps' m j ks2) s" + thus "reps' i j (ks1 @ ks2) s" + proof(elim EXS_elim) + fix m + assume "(reps' i (m - 1) ks1 \* reps' m j ks2) s" + then obtain s1 s2 where h: + "s = s1 + s2" "s1 ## s2" "reps' i (m - 1) ks1 s1" + " reps' m j ks2 s2" by (auto elim!:sep_conjE) + from reps'_len_correct[OF this(3) `ks1 \ []`] + have eq_m: "m = i + int (reps_len ks1) + 1" by simp + have "((reps i (i + int (reps_len ks1) - 1) ks1 \* zero (i + int (reps_len ks1))) \* + ((reps (i + int (reps_len ks1) + 1) (j - 1) ks2) \* zero j)) s" + (is "(?P \* ?Q) s") + proof(rule sep_conjI) + from h(3) eq_m `ks1 \ []` show "?P s1" + by (simp add:reps'_def) + next + from h(4) eq_m `ks2 \ []` show "?Q s2" + by (simp add:reps'_def) + next + from h(2) show "s1 ## s2" . + next + from h(1) show "s = s1 + s2" . + qed + thus "reps' i j (ks1 @ ks2) s" + by (unfold eq_2 eq_1, auto simp:sep_conj_ac) + qed + qed + next + case True + have "-1 + j = j - 1" by auto + with `ks1 \ []` True + show ?thesis + apply (auto simp:reps'_def pred_ex_def pasrt_def) + apply (unfold `-1 + j = j - 1`) + by (fold sep_empty_def, simp only:sep_conj_empty) + qed +qed + +lemma reps'_sg: + "reps' i j [k] = + (<(j = i + int k + 1)> \* ones i (i + int k) \* zero j)" + apply (unfold reps'_def, default, auto simp:sep_conj_ac) + by (sep_cancel+, simp add:pasrt_def)+ + + +section {* Basic macros for TM *} + +lemma st_upd: + assumes pre: "(st i' ** r) (trset_of (f, x, i, y, z))" + shows "(st i'' ** r) (trset_of (f, x, i'', y, z))" +proof - + from stimes_sgD[OF pre[unfolded st_def], unfolded trset_of.simps, unfolded stD[OF pre]] + have "r (tprog_set x \ tpc_set i' \ tpos_set y \ tmem_set z \ tfaults_set f - tpc_set i')" + by blast + moreover have + "(tprog_set x \ tpc_set i' \ tpos_set y \ tmem_set z \ tfaults_set f - tpc_set i') = + (tprog_set x \ tpos_set y \ tmem_set z \ tfaults_set f)" + by (unfold tpn_set_def, auto) + ultimately have r_rest: "r (tprog_set x \ tpos_set y \ tmem_set z \ tfaults_set f)" + by auto + show ?thesis + proof(rule sep_conjI[where Q = r, OF _ r_rest]) + show "st i'' (tpc_set i'')" + by (unfold st_def sg_def, simp) + next + show "tpc_set i'' ## tprog_set x \ tpos_set y \ tmem_set z \ tfaults_set f" + by (unfold tpn_set_def sep_disj_set_def, auto) + next + show "trset_of (f, x, i'', y, z) = + tpc_set i'' + (tprog_set x \ tpos_set y \ tmem_set z \ tfaults_set f)" + by (unfold trset_of.simps plus_set_def, auto) + qed +qed + +lemma pos_upd: + assumes pre: "(ps i ** r) (trset_of (f, x, y, i', z))" + shows "(ps j ** r) (trset_of (f, x, y, j, z))" +proof - + from stimes_sgD[OF pre[unfolded ps_def], unfolded trset_of.simps, unfolded psD[OF pre]] + have "r (tprog_set x \ tpc_set y \ tpos_set i \ tmem_set z \ + tfaults_set f - tpos_set i)" (is "r ?xs") + by blast + moreover have + "?xs = tprog_set x \ tpc_set y \ tmem_set z \ tfaults_set f" + by (unfold tpn_set_def, auto) + ultimately have r_rest: "r \" + by auto + show ?thesis + proof(rule sep_conjI[where Q = r, OF _ r_rest]) + show "ps j (tpos_set j)" + by (unfold ps_def sg_def, simp) + next + show "tpos_set j ## tprog_set x \ tpc_set y \ tmem_set z \ tfaults_set f" + by (unfold tpn_set_def sep_disj_set_def, auto) + next + show "trset_of (f, x, y, j, z) = + tpos_set j + (tprog_set x \ tpc_set y \ tmem_set z \ tfaults_set f)" + by (unfold trset_of.simps plus_set_def, auto) + qed +qed + +lemma TM_in_simp: "({TM a v} \ + tprog_set x \ tpc_set y \ tpos_set z \ tmem_set mem \ tfaults_set f) = + ({TM a v} \ tmem_set mem)" + by (unfold tpn_set_def, auto) + +lemma tmem_set_upd: + "{TM a v} \ tmem_set mem \ + tmem_set (mem(a:=Some v')) = ((tmem_set mem) - {TM a v}) \ {TM a v'}" + by (unfold tpn_set_def, auto) + +lemma tmem_set_disj: "{TM a v} \ tmem_set mem \ + {TM a v'} \ (tmem_set mem - {TM a v}) = {}" + by (unfold tpn_set_def, auto) + +lemma smem_upd: "((tm a v) ** r) (trset_of (f, x, y, z, mem)) \ + ((tm a v') ** r) (trset_of (f, x, y, z, mem(a := Some v')))" +proof - + have eq_s: "(tprog_set x \ tpc_set y \ tpos_set z \ tmem_set mem \ tfaults_set f - {TM a v}) = + (tprog_set x \ tpc_set y \ tpos_set z \ (tmem_set mem - {TM a v}) \ tfaults_set f)" + by (unfold tpn_set_def, auto) + assume g: "(tm a v \* r) (trset_of (f, x, y, z, mem))" + from this[unfolded trset_of.simps tm_def] + have h: "(sg {TM a v} \* r) (tprog_set x \ tpc_set y \ tpos_set z \ tmem_set mem \ tfaults_set f)" . + hence h0: "r (tprog_set x \ tpc_set y \ tpos_set z \ (tmem_set mem - {TM a v}) \ tfaults_set f)" + by(sep_drule stimes_sgD, clarify, unfold eq_s, auto) + from h TM_in_simp have "{TM a v} \ tmem_set mem" + by(sep_drule stimes_sgD, auto) + from tmem_set_upd [OF this] tmem_set_disj[OF this] + have h2: "tmem_set (mem(a \ v')) = {TM a v'} \ (tmem_set mem - {TM a v})" + "{TM a v'} \ (tmem_set mem - {TM a v}) = {}" by auto + show ?thesis + proof - + have "(tm a v' ** r) (tmem_set (mem(a \ v')) \ tprog_set x \ tpc_set y \ tpos_set z \ tfaults_set f)" + proof(rule sep_conjI) + show "tm a v' ({TM a v'})" by (unfold tm_def sg_def, simp) + next + from h0 show "r (tprog_set x \ tpc_set y \ tpos_set z \ (tmem_set mem - {TM a v}) \ tfaults_set f)" . + next + show "{TM a v'} ## tprog_set x \ tpc_set y \ tpos_set z \ (tmem_set mem - {TM a v}) \ tfaults_set f" + proof - + from g have " mem a = Some v" + by(sep_frule memD, simp) + thus "?thesis" + by(unfold tpn_set_def set_ins_def, auto) + qed + next + show "tmem_set (mem(a \ v')) \ tprog_set x \ tpc_set y \ tpos_set z \ tfaults_set f = + {TM a v'} + (tprog_set x \ tpc_set y \ tpos_set z \ (tmem_set mem - {TM a v}) \ tfaults_set f)" + by (unfold h2(1) set_ins_def eq_s, auto) + qed + thus ?thesis + apply (unfold trset_of.simps) + by (metis sup_commute sup_left_commute) + qed +qed + +lemma hoare_write_zero: + "\st i ** ps p ** tm p v\ + i:[write_zero]:j + \st j ** ps p ** tm p Bk\" +proof(unfold write_zero_def, intro t_hoare_local, rule t_hoare_label_last, simp, simp) + show "\st i \* ps p \* tm p v\ i :[ \ ((W0, j), W0, j) ]: j \st j \* ps p \* tm p Bk\" + proof(unfold tassemble_to.simps, simp only:sep_conj_cond, + intro tm.code_condI, simp) + assume eq_j: "j = Suc i" + show "\st i \* ps p \* tm p v\ sg {TC i ((W0, Suc i), W0, Suc i)} + \st (Suc i) \* ps p \* tm p Bk\" + proof(fold eq_j, unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) + fix ft prog cs i' mem r + assume h: "(r \* ps p \* st i \* tm p v \* sg {TC i ((W0, j), W0, j)}) + (trset_of (ft, prog, cs, i', mem))" + from h have "prog i = Some ((W0, j), W0, j)" + apply(rule_tac r = "r \* ps p \* tm p v" in codeD) + by(simp add: sep_conj_ac) + from h and this have stp: + "tm.run 1 (ft, prog, cs, i', mem) = (ft, prog, j, i', mem(i'\ Bk))" (is "?x = ?y") + apply(sep_frule psD) + apply(sep_frule stD) + apply(sep_frule memD, simp) + by(cases v, unfold tm.run_def, auto) + from h have "i' = p" + by(sep_drule psD, simp) + with h + have "(r \* ps p \* st j \* tm p Bk \* sg {TC i ((W0, j), W0, j)}) (trset_of ?x)" + apply(unfold stp) + apply(sep_drule pos_upd, sep_drule st_upd, sep_drule smem_upd) + apply(auto simp: sep_conj_ac) + done + thus "\k. (r \* ps p \* st j \* tm p Bk \* sg {TC i ((W0, j), W0, j)}) + (trset_of (tm.run (Suc k) (ft, prog, cs, i', mem)))" + apply (rule_tac x = 0 in exI) + by auto + qed + qed +qed + +lemma hoare_write_zero_gen[step]: + assumes "p = q" + shows "\st i ** ps p ** tm q v\ + i:[write_zero]:j + \st j ** ps p ** tm q Bk\" + by (unfold assms, rule hoare_write_zero) + +lemma hoare_write_one: + "\st i ** ps p ** tm p v\ + i:[write_one]:j + \st j ** ps p ** tm p Oc\" +proof(unfold write_one_def, intro t_hoare_local, rule t_hoare_label_last, simp+) + fix l + show "\st i \* ps p \* tm p v\ i :[ \ ((W1, j), W1, j) ]: j \st j \* ps p \* tm p Oc\" + proof(unfold tassemble_to.simps, simp only:sep_conj_cond, + rule_tac tm.code_condI, simp add: sep_conj_ac) + let ?j = "Suc i" + show "\ps p \* st i \* tm p v\ sg {TC i ((W1, ?j), W1, ?j)} + \ps p \* st ?j \* tm p Oc\" + proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) + fix ft prog cs i' mem r + assume h: "(r \* ps p \* st i \* tm p v \* sg {TC i ((W1, ?j), W1, ?j)}) + (trset_of (ft, prog, cs, i', mem))" + from h have "prog i = Some ((W1, ?j), W1, ?j)" + apply(rule_tac r = "r \* ps p \* tm p v" in codeD) + by(simp add: sep_conj_ac) + from h and this have stp: + "tm.run 1 (ft, prog, cs, i', mem) = + (ft, prog, ?j, i', mem(i'\ Oc))" (is "?x = ?y") + apply(sep_frule psD) + apply(sep_frule stD) + apply(sep_frule memD, simp) + by(cases v, unfold tm.run_def, auto) + from h have "i' = p" + by(sep_drule psD, simp) + with h + have "(r \* ps p \* st ?j \* tm p Oc \* sg {TC i ((W1, ?j), W1, ?j)}) (trset_of ?x)" + apply(unfold stp) + apply(sep_drule pos_upd, sep_drule st_upd, sep_drule smem_upd) + apply(auto simp: sep_conj_ac) + done + thus "\k. (r \* ps p \* st ?j \* tm p Oc \* sg {TC i ((W1, ?j), W1, ?j)}) + (trset_of (tm.run (Suc k) (ft, prog, cs, i', mem)))" + apply (rule_tac x = 0 in exI) + by auto + qed + qed +qed + +lemma hoare_write_one_gen[step]: + assumes "p = q" + shows "\st i ** ps p ** tm q v\ + i:[write_one]:j + \st j ** ps p ** tm q Oc\" + by (unfold assms, rule hoare_write_one) + +lemma hoare_move_left: + "\st i ** ps p ** tm p v2\ + i:[move_left]:j + \st j ** ps (p - 1) ** tm p v2\" +proof(unfold move_left_def, intro t_hoare_local, rule t_hoare_label_last, simp+) + fix l + show "\st i \* ps p \* tm p v2\ i :[ \ ((L, l), L, l) ]: l + \st l \* ps (p - 1) \* tm p v2\" + proof(unfold tassemble_to.simps, simp only:sep_conj_cond, + intro tm.code_condI, simp add: sep_conj_ac) + let ?j = "Suc i" + show "\ps p \* st i \* tm p v2\ sg {TC i ((L, ?j), L, ?j)} + \st ?j \* tm p v2 \* ps (p - 1)\" + proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) + fix ft prog cs i' mem r + assume h: "(r \* ps p \* st i \* tm p v2 \* sg {TC i ((L, ?j), L, ?j)}) + (trset_of (ft, prog, cs, i', mem))" + from h have "prog i = Some ((L, ?j), L, ?j)" + apply(rule_tac r = "r \* ps p \* tm p v2" in codeD) + by(simp add: sep_conj_ac) + from h and this have stp: + "tm.run 1 (ft, prog, cs, i', mem) = (ft, prog, ?j, i' - 1, mem)" (is "?x = ?y") + apply(sep_frule psD) + apply(sep_frule stD) + apply(sep_frule memD, simp) + apply(unfold tm.run_def, case_tac v2, auto) + done + from h have "i' = p" + by(sep_drule psD, simp) + with h + have "(r \* st ?j \* tm p v2 \* ps (p - 1) \* sg {TC i ((L, ?j), L, ?j)}) + (trset_of ?x)" + apply(unfold stp) + apply(sep_drule pos_upd, sep_drule st_upd, simp) + proof - + assume "(st ?j \* ps (p - 1) \* r \* tm p v2 \* sg {TC i ((L, ?j), L, ?j)}) + (trset_of (ft, prog, ?j, p - 1, mem))" + thus "(r \* st ?j \* tm p v2 \* ps (p - 1) \* sg {TC i ((L, ?j), L, ?j)}) + (trset_of (ft, prog, ?j, p - 1, mem))" + by(simp add: sep_conj_ac) + qed + thus "\k. (r \* st ?j \* tm p v2 \* ps (p - 1) \* sg {TC i ((L, ?j), L, ?j)}) + (trset_of (tm.run (Suc k) (ft, prog, cs, i', mem)))" + apply (rule_tac x = 0 in exI) + by auto + qed + qed +qed + +lemma hoare_move_left_gen[step]: + assumes "p = q" + shows "\st i ** ps p ** tm q v2\ + i:[move_left]:j + \st j ** ps (p - 1) ** tm q v2\" + by (unfold assms, rule hoare_move_left) + +lemma hoare_move_right: + "\st i ** ps p ** tm p v1 \ + i:[move_right]:j + \st j ** ps (p + 1) ** tm p v1 \" +proof(unfold move_right_def, intro t_hoare_local, rule t_hoare_label_last, simp+) + fix l + show "\st i \* ps p \* tm p v1\ i :[ \ ((R, l), R, l) ]: l + \st l \* ps (p + 1) \* tm p v1\" + proof(unfold tassemble_to.simps, simp only:sep_conj_cond, + intro tm.code_condI, simp add: sep_conj_ac) + let ?j = "Suc i" + show "\ps p \* st i \* tm p v1\ sg {TC i ((R, ?j), R, ?j)} + \st ?j \* tm p v1 \* ps (p + 1)\" + proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) + fix ft prog cs i' mem r + assume h: "(r \* ps p \* st i \* tm p v1 \* sg {TC i ((R, ?j), R, ?j)}) + (trset_of (ft, prog, cs, i', mem))" + from h have "prog i = Some ((R, ?j), R, ?j)" + apply(rule_tac r = "r \* ps p \* tm p v1" in codeD) + by(simp add: sep_conj_ac) + from h and this have stp: + "tm.run 1 (ft, prog, cs, i', mem) = (ft, prog, ?j, i'+ 1, mem)" (is "?x = ?y") + apply(sep_frule psD) + apply(sep_frule stD) + apply(sep_frule memD, simp) + apply(unfold tm.run_def, case_tac v1, auto) + done + from h have "i' = p" + by(sep_drule psD, simp) + with h + have "(r \* st ?j \* tm p v1 \* ps (p + 1) \* + sg {TC i ((R, ?j), R, ?j)}) (trset_of ?x)" + apply(unfold stp) + apply(sep_drule pos_upd, sep_drule st_upd, simp) + proof - + assume "(st ?j \* ps (p + 1) \* r \* tm p v1 \* sg {TC i ((R, ?j), R, ?j)}) + (trset_of (ft, prog, ?j, p + 1, mem))" + thus "(r \* st ?j \* tm p v1 \* ps (p + 1) \* sg {TC i ((R, ?j), R, ?j)}) + (trset_of (ft, prog, ?j, p + 1, mem))" + by (simp add: sep_conj_ac) + qed + thus "\k. (r \* st ?j \* tm p v1 \* ps (p + 1) \* sg {TC i ((R, ?j), R, ?j)}) + (trset_of (tm.run (Suc k) (ft, prog, cs, i', mem)))" + apply (rule_tac x = 0 in exI) + by auto + qed + qed +qed + +lemma hoare_move_right_gen[step]: + assumes "p = q" + shows "\st i ** ps p ** tm q v1 \ + i:[move_right]:j + \st j ** ps (p + 1) ** tm q v1 \" + by (unfold assms, rule hoare_move_right) + +lemma hoare_if_one_true: + "\st i ** ps p ** one p\ + i:[if_one e]:j + \st e ** ps p ** one p\" +proof(unfold if_one_def, intro t_hoare_local, rule t_hoare_label_last, simp+) + fix l + show " \st i \* ps p \* one p\ i :[ \ ((W0, l), W1, e) ]: l \st e \* ps p \* one p\" + proof(unfold tassemble_to.simps, simp only:sep_conj_cond, + intro tm.code_condI, simp add: sep_conj_ac) + let ?j = "Suc i" + show "\one p \* ps p \* st i\ sg {TC i ((W0, ?j), W1, e)} \one p \* ps p \* st e\" + proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) + fix ft prog cs pc mem r + assume h: "(r \* one p \* ps p \* st i \* sg {TC i ((W0, ?j), W1, e)}) + (trset_of (ft, prog, cs, pc, mem))" + from h have k1: "prog i = Some ((W0, ?j), W1, e)" + apply(rule_tac r = "r \* one p \* ps p" in codeD) + by(simp add: sep_conj_ac) + from h have k2: "pc = p" + by(sep_drule psD, simp) + from h and this have k3: "mem pc = Some Oc" + apply(unfold one_def) + by(sep_drule memD, simp) + from h k1 k2 k3 have stp: + "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, e, pc, mem)" (is "?x = ?y") + apply(sep_drule stD) + by(unfold tm.run_def, auto) + from h k2 + have "(r \* one p \* ps p \* st e \* sg {TC i ((W0, ?j), W1, e)}) (trset_of ?x)" + apply(unfold stp) + by(sep_drule st_upd, simp add: sep_conj_ac) + thus "\k.(r \* one p \* ps p \* st e \* sg {TC i ((W0, ?j), W1, e)}) + (trset_of (tm.run (Suc k) (ft, prog, cs, pc, mem)))" + apply (rule_tac x = 0 in exI) + by auto + qed + qed +qed + +text {* + The following problematic lemma is not provable now + lemma hoare_self: "\p\ i :[ap]: j \p\" + apply(simp add: tm.Hoare_gen_def, clarify) + apply(rule_tac x = 0 in exI, simp add: tm.run_def) + done +*} + +lemma hoare_if_one_true_gen[step]: + assumes "p = q" + shows + "\st i ** ps p ** one q\ + i:[if_one e]:j + \st e ** ps p ** one q\" + by (unfold assms, rule hoare_if_one_true) + +lemma hoare_if_one_true1: + "\st i ** ps p ** one p\ + i:[(if_one e; c)]:j + \st e ** ps p ** one p\" +proof(unfold tassemble_to.simps, rule tm.code_exI, + simp add: sep_conj_ac tm.Hoare_gen_def, clarify) + fix j' ft prog cs pos mem r + assume h: "(r \* one p \* ps p \* st i \* j' :[ c ]: j \* i :[ if_one e ]: j') + (trset_of (ft, prog, cs, pos, mem))" + from tm.frame_rule[OF hoare_if_one_true] + have "\ r. \st i \* ps p \* one p \* r\ i :[ if_one e ]: j' \st e \* ps p \* one p \* r\" + by(simp add: sep_conj_ac) + from this[unfolded tm.Hoare_gen_def tassemble_to.simps, rule_format, of "j' :[ c ]: j"] h + have "\ k. (r \* one p \* ps p \* st e \* i :[ if_one e ]: j' \* j' :[ c ]: j) + (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))" + by(auto simp: sep_conj_ac) + thus "\k. (r \* one p \* ps p \* st e \* j' :[ c ]: j \* i :[ if_one e ]: j') + (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))" + by(simp add: sep_conj_ac) +qed + +lemma hoare_if_one_true1_gen[step]: + assumes "p = q" + shows + "\st i ** ps p ** one q\ + i:[(if_one e; c)]:j + \st e ** ps p ** one q\" + by (unfold assms, rule hoare_if_one_true1) + +lemma hoare_if_one_false: + "\st i ** ps p ** zero p\ + i:[if_one e]:j + \st j ** ps p ** zero p\" +proof(unfold if_one_def, intro t_hoare_local, rule t_hoare_label_last, simp+) + show "\st i \* ps p \* zero p\ i :[ (\ ((W0, j), W1, e)) ]: j + \st j \* ps p \* zero p\" + proof(unfold tassemble_to.simps, simp only:sep_conj_cond, + intro tm.code_condI, simp add: sep_conj_ac) + let ?j = "Suc i" + show "\ps p \* st i \* zero p\ sg {TC i ((W0, ?j), W1, e)} \ps p \* zero p \* st ?j \" + proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) + fix ft prog cs pc mem r + assume h: "(r \* ps p \* st i \* zero p \* sg {TC i ((W0, ?j), W1, e)}) + (trset_of (ft, prog, cs, pc, mem))" + from h have k1: "prog i = Some ((W0, ?j), W1, e)" + apply(rule_tac r = "r \* zero p \* ps p" in codeD) + by(simp add: sep_conj_ac) + from h have k2: "pc = p" + by(sep_drule psD, simp) + from h and this have k3: "mem pc = Some Bk" + apply(unfold zero_def) + by(sep_drule memD, simp) + from h k1 k2 k3 have stp: + "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, ?j, pc, mem)" (is "?x = ?y") + apply(sep_drule stD) + by (unfold tm.run_def, auto split:tstate.splits) + from h k2 + have "(r \* zero p \* ps p \* st ?j \* sg {TC i ((W0, ?j), W1, e)}) (trset_of ?x)" + apply (unfold stp) + by (sep_drule st_upd[where i''="?j"], auto simp:sep_conj_ac) + thus "\k. (r \* ps p \* zero p \* st ?j \* sg {TC i ((W0, ?j), W1, e)}) + (trset_of (tm.run (Suc k) (ft, prog, cs, pc, mem)))" + by(auto simp: sep_conj_ac) + qed + qed +qed + +lemma hoare_if_one_false_gen[step]: + assumes "p = q" + shows "\st i ** ps p ** zero q\ + i:[if_one e]:j + \st j ** ps p ** zero q\" + by (unfold assms, rule hoare_if_one_false) + +lemma hoare_if_zero_true: + "\st i ** ps p ** zero p\ + i:[if_zero e]:j + \st e ** ps p ** zero p\" +proof(unfold if_zero_def, intro t_hoare_local, rule t_hoare_label_last, simp+) + fix l + show "\st i \* ps p \* zero p\ i :[ \ ((W0, e), W1, l) ]: l \st e \* ps p \* zero p\" + proof(unfold tassemble_to.simps, simp only:sep_conj_cond, + intro tm.code_condI, simp add: sep_conj_ac) + let ?j = "Suc i" + show "\ps p \* st i \* zero p\ sg {TC i ((W0, e), W1, ?j)} \ps p \* st e \* zero p\" + proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) + fix ft prog cs pc mem r + assume h: "(r \* ps p \* st i \* zero p \* sg {TC i ((W0, e), W1, ?j)}) + (trset_of (ft, prog, cs, pc, mem))" + from h have k1: "prog i = Some ((W0, e), W1, ?j)" + apply(rule_tac r = "r \* zero p \* ps p" in codeD) + by(simp add: sep_conj_ac) + from h have k2: "pc = p" + by(sep_drule psD, simp) + from h and this have k3: "mem pc = Some Bk" + apply(unfold zero_def) + by(sep_drule memD, simp) + from h k1 k2 k3 have stp: + "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, e, pc, mem)" (is "?x = ?y") + apply(sep_drule stD) + by(unfold tm.run_def, auto) + from h k2 + have "(r \* zero p \* ps p \* st e \* sg {TC i ((W0, e), W1, ?j)}) (trset_of ?x)" + apply(unfold stp) + by(sep_drule st_upd, simp add: sep_conj_ac) + thus "\k. (r \* ps p \* st e \* zero p \* sg {TC i ((W0, e), W1, ?j)}) + (trset_of (tm.run (Suc k) (ft, prog, cs, pc, mem)))" + by(auto simp: sep_conj_ac) + qed + qed +qed + +lemma hoare_if_zero_true_gen[step]: + assumes "p = q" + shows + "\st i ** ps p ** zero q\ + i:[if_zero e]:j + \st e ** ps p ** zero q\" + by (unfold assms, rule hoare_if_zero_true) + +lemma hoare_if_zero_true1: + "\st i ** ps p ** zero p\ + i:[(if_zero e; c)]:j + \st e ** ps p ** zero p\" + proof(unfold tassemble_to.simps, rule tm.code_exI, simp add: sep_conj_ac + tm.Hoare_gen_def, clarify) + fix j' ft prog cs pos mem r + assume h: "(r \* ps p \* st i \* zero p \* j' :[ c ]: j \* i :[ if_zero e ]: j') + (trset_of (ft, prog, cs, pos, mem))" + from tm.frame_rule[OF hoare_if_zero_true] + have "\ r. \st i \* ps p \* zero p \* r\ i :[ if_zero e ]: j' \st e \* ps p \* zero p \* r\" + by(simp add: sep_conj_ac) + from this[unfolded tm.Hoare_gen_def tassemble_to.simps, rule_format, of "j' :[ c ]: j"] h + have "\ k. (r \* zero p \* ps p \* st e \* i :[ if_zero e ]: j' \* j' :[ c ]: j) + (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))" + by(auto simp: sep_conj_ac) + thus "\k. (r \* ps p \* st e \* zero p \* j' :[ c ]: j \* i :[ if_zero e ]: j') + (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))" + by(simp add: sep_conj_ac) +qed + +lemma hoare_if_zero_true1_gen[step]: + assumes "p = q" + shows + "\st i ** ps p ** zero q\ + i:[(if_zero e; c)]:j + \st e ** ps p ** zero q\" + by (unfold assms, rule hoare_if_zero_true1) + +lemma hoare_if_zero_false: + "\st i ** ps p ** tm p Oc\ + i:[if_zero e]:j + \st j ** ps p ** tm p Oc\" +proof(unfold if_zero_def, intro t_hoare_local, rule t_hoare_label_last, simp, simp) + fix l + show "\st i \* ps p \* tm p Oc\ i :[ \ ((W0, e), W1, l) ]: l + \st l \* ps p \* tm p Oc\" + proof(unfold tassemble_to.simps, simp only:sep_conj_cond, + intro tm.code_condI, simp add: sep_conj_ac) + let ?j = "Suc i" + show "\ps p \* st i \* tm p Oc\ sg {TC i ((W0, e), W1, ?j)} + \ps p \* st ?j \* tm p Oc\" + proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) + fix ft prog cs pc mem r + assume h: "(r \* ps p \* st i \* tm p Oc \* sg {TC i ((W0, e), W1, ?j)}) + (trset_of (ft, prog, cs, pc, mem))" + from h have k1: "prog i = Some ((W0, e), W1, ?j)" + apply(rule_tac r = "r \* tm p Oc \* ps p" in codeD) + by(simp add: sep_conj_ac) + from h have k2: "pc = p" + by(sep_drule psD, simp) + from h and this have k3: "mem pc = Some Oc" + by(sep_drule memD, simp) + from h k1 k2 k3 have stp: + "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, ?j, pc, mem)" (is "?x = ?y") + apply(sep_drule stD) + by(unfold tm.run_def, auto split:tstate.splits) + from h k2 + have "(r \* tm p Oc \* ps p \* st ?j \* sg {TC i ((W0, e), W1, ?j)}) (trset_of ?x)" + apply(unfold stp) + by(sep_drule st_upd, simp add: sep_conj_ac) + thus "\k. (r \* ps p \* st ?j \* tm p Oc \* sg {TC i ((W0, e), W1, ?j)}) + (trset_of (tm.run (Suc k) (ft, prog, cs, pc, mem)))" + by(auto simp: sep_conj_ac) + qed + qed +qed + +lemma hoare_if_zero_false_gen[step]: + assumes "p = q" + shows + "\st i ** ps p ** tm q Oc\ + i:[if_zero e]:j + \st j ** ps p ** tm q Oc\" + by (unfold assms, rule hoare_if_zero_false) + +lemma hoare_jmp: + "\st i \* ps p \* tm p v\ i:[jmp e]:j \st e \* ps p \* tm p v\" +proof(unfold jmp_def tm.Hoare_gen_def tassemble_to.simps sep_conj_ac, clarify) + fix ft prog cs pos mem r + assume h: "(r \* ps p \* st i \* tm p v \* <(j = i + 1)> \* sg {TC i ((W0, e), W1, e)}) + (trset_of (ft, prog, cs, pos, mem))" + from h have k1: "prog i = Some ((W0, e), W1, e)" + apply(rule_tac r = "r \* <(j = i + 1)> \* tm p v \* ps p" in codeD) + by(simp add: sep_conj_ac) + from h have k2: "p = pos" + by(sep_drule psD, simp) + from h k2 have k3: "mem pos = Some v" + by(sep_drule memD, simp) + from h k1 k2 k3 have + stp: "tm.run 1 (ft, prog, cs, pos, mem) = (ft, prog, e, pos, mem)" (is "?x = ?y") + apply(sep_drule stD) + by(unfold tm.run_def, cases "mem pos", simp, cases v, auto) + from h k2 + have "(r \* ps p \* st e \* tm p v \* <(j = i + 1)> \* + sg {TC i ((W0, e), W1, e)}) (trset_of ?x)" + apply(unfold stp) + by(sep_drule st_upd, simp add: sep_conj_ac) + thus "\ k. (r \* ps p \* st e \* tm p v \* <(j = i + 1)> \* sg {TC i ((W0, e), W1, e)}) + (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))" + apply (rule_tac x = 0 in exI) + by auto +qed + +lemma hoare_jmp_gen[step]: + assumes "p = q" + shows "\st i \* ps p \* tm q v\ i:[jmp e]:j \st e \* ps p \* tm q v\" + by (unfold assms, rule hoare_jmp) + +lemma hoare_jmp1: + "\st i \* ps p \* tm p v\ + i:[(jmp e; c)]:j + \st e \* ps p \* tm p v\" +proof(unfold tassemble_to.simps, rule tm.code_exI, simp + add: sep_conj_ac tm.Hoare_gen_def, clarify) + fix j' ft prog cs pos mem r + assume h: "(r \* ps p \* st i \* tm p v \* j' :[ c ]: j \* i :[ jmp e ]: j') + (trset_of (ft, prog, cs, pos, mem))" + from tm.frame_rule[OF hoare_jmp] + have "\ r. \st i \* ps p \* tm p v \* r\ i :[ jmp e ]: j' \st e \* ps p \* tm p v \* r\" + by(simp add: sep_conj_ac) + from this[unfolded tm.Hoare_gen_def tassemble_to.simps, rule_format, of "j' :[ c ]: j"] h + have "\ k. (r \* tm p v \* ps p \* st e \* i :[ jmp e ]: j' \* j' :[ c ]: j) + (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))" + by(auto simp: sep_conj_ac) + thus "\k. (r \* ps p \* st e \* tm p v \* j' :[ c ]: j \* i :[ jmp e ]: j') + (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))" + by(simp add: sep_conj_ac) +qed + + +lemma hoare_jmp1_gen[step]: + assumes "p = q" + shows "\st i \* ps p \* tm q v\ + i:[(jmp e; c)]:j + \st e \* ps p \* tm q v\" + by (unfold assms, rule hoare_jmp1) + + +lemma condI: + assumes h1: b + and h2: "b \ p s" + shows "( \* p) s" + by (metis (full_types) cond_true_eq1 h1 h2) + +lemma condE: + assumes "( \* p) s" + obtains "b" and "p s" +proof(atomize_elim) + from condD[OF assms] + show "b \ p s" . +qed + + +section {* Tactics *} + +ML {* + val trace_step = Attrib.setup_config_bool @{binding trace_step} (K false) + val trace_fwd = Attrib.setup_config_bool @{binding trace_fwd} (K false) +*} + + +ML {* + val tracing = (fn ctxt => fn str => + if (Config.get ctxt trace_step) then tracing str else ()) + fun not_pred p = fn s => not (p s) + fun break_sep_conj (Const (@{const_name sep_conj},_) $ t1 $ t2 $ _) = + (break_sep_conj t1) @ (break_sep_conj t2) + | break_sep_conj (Const (@{const_name sep_conj},_) $ t1 $ t2) = + (break_sep_conj t1) @ (break_sep_conj t2) + (* dig through eta exanded terms: *) + | break_sep_conj (Abs (_, _, t $ Bound 0)) = break_sep_conj t + | break_sep_conj t = [t]; + + val empty_env = (Vartab.empty, Vartab.empty) + + fun match_env ctxt pat trm env = + Pattern.match (ctxt |> Proof_Context.theory_of) (pat, trm) env + + fun match ctxt pat trm = match_env ctxt pat trm empty_env; + + val inst = Envir.subst_term; + + fun term_of_thm thm = thm |> prop_of |> HOLogic.dest_Trueprop + + fun get_cmd ctxt code = + let val pat = term_of @{cpat "_:[(?cmd)]:_"} + val pat1 = term_of @{cpat "?cmd::tpg"} + val env = match ctxt pat code + in inst env pat1 end + + fun is_seq_term (Const (@{const_name TSeq}, _) $ _ $ _) = true + | is_seq_term _ = false + + fun get_hcmd (Const (@{const_name TSeq}, _) $ hcmd $ _) = hcmd + | get_hcmd hcmd = hcmd + + fun last [a] = a | + last (a::b) = last b + + fun but_last [a] = [] | + but_last (a::b) = a::(but_last b) + + fun foldr f [] = (fn x => x) | + foldr f (x :: xs) = (f x) o (foldr f xs) + + fun concat [] = [] | + concat (x :: xs) = x @ concat xs + + fun match_any ctxt pats tm = + fold + (fn pat => fn b => (b orelse Pattern.matches + (ctxt |> Proof_Context.theory_of) (pat, tm))) + pats false + + fun is_ps_term (Const (@{const_name ps}, _) $ _) = true + | is_ps_term _ = false + + fun string_of_term ctxt t = t |> Syntax.pretty_term ctxt |> Pretty.str_of + fun string_of_cterm ctxt ct = ct |> term_of |> string_of_term ctxt + fun pterm ctxt t = + t |> string_of_term ctxt |> tracing ctxt + fun pcterm ctxt ct = ct |> string_of_cterm ctxt |> tracing ctxt + fun string_for_term ctxt t = + Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN) + (print_mode_value ())) (Syntax.string_of_term ctxt) t + |> String.translate (fn c => if Char.isPrint c then str c else "") + |> Sledgehammer_Util.simplify_spaces + fun string_for_cterm ctxt ct = ct |> term_of |> string_for_term ctxt + fun attemp tac = fn i => fn st => (tac i st) handle exn => Seq.empty + fun try_tac tac = fn i => fn st => (tac i st) handle exn => (Seq.single st) + (* aux end *) +*} + +ML {* (* Functions specific to Hoare triples *) + fun get_pre ctxt t = + let val pat = term_of @{cpat "\?P\ ?c \?Q\"} + val env = match ctxt pat t + in inst env (term_of @{cpat "?P::tresource set \ bool"}) end + + fun can_process ctxt t = ((get_pre ctxt t; true) handle _ => false) + + fun get_post ctxt t = + let val pat = term_of @{cpat "\?P\ ?c \?Q\"} + val env = match ctxt pat t + in inst env (term_of @{cpat "?Q::tresource set \ bool"}) end; + + fun get_mid ctxt t = + let val pat = term_of @{cpat "\?P\ ?c \?Q\"} + val env = match ctxt pat t + in inst env (term_of @{cpat "?c::tresource set \ bool"}) end; + + fun is_pc_term (Const (@{const_name st}, _) $ _) = true + | is_pc_term _ = false + + fun mk_pc_term x = + Const (@{const_name st}, @{typ "nat \ tresource set \ bool"}) $ Free (x, @{typ "nat"}) + + val sconj_term = term_of @{cterm "sep_conj::tassert \ tassert \ tassert"} + + fun mk_ps_term x = + Const (@{const_name ps}, @{typ "int \ tresource set \ bool"}) $ Free (x, @{typ "int"}) + + fun atomic tac = ((SOLVED' tac) ORELSE' (K all_tac)) + + fun pure_sep_conj_ac_tac ctxt = + (auto_tac (ctxt |> Simplifier.map_simpset (fn ss => ss addsimps @{thms sep_conj_ac})) + |> SELECT_GOAL) + + + fun potential_facts ctxt prop = Facts.could_unify (Proof_Context.facts_of ctxt) + ((Term.strip_all_body prop) |> Logic.strip_imp_concl); + + fun some_fact_tac ctxt = SUBGOAL (fn (goal, i) => + (Method.insert_tac (potential_facts ctxt goal) i) THEN + (pure_sep_conj_ac_tac ctxt i)); + + fun sep_conj_ac_tac ctxt = + (SOLVED' (auto_tac (ctxt |> Simplifier.map_simpset (fn ss => ss addsimps @{thms sep_conj_ac})) + |> SELECT_GOAL)) ORELSE' (atomic (some_fact_tac ctxt)) +*} + +ML {* +type HoareTriple = { + binding: binding, + can_process: Proof.context -> term -> bool, + get_pre: Proof.context -> term -> term, + get_mid: Proof.context -> term -> term, + get_post: Proof.context -> term -> term, + is_pc_term: term -> bool, + mk_pc_term: string -> term, + sconj_term: term, + sep_conj_ac_tac: Proof.context -> int -> tactic, + hoare_seq1: thm, + hoare_seq2: thm, + pre_stren: thm, + post_weaken: thm, + frame_rule: thm +} + + val tm_triple = {binding = @{binding "tm_triple"}, + can_process = can_process, + get_pre = get_pre, + get_mid = get_mid, + get_post = get_post, + is_pc_term = is_pc_term, + mk_pc_term = mk_pc_term, + sconj_term = sconj_term, + sep_conj_ac_tac = sep_conj_ac_tac, + hoare_seq1 = @{thm t_hoare_seq1}, + hoare_seq2 = @{thm t_hoare_seq2}, + pre_stren = @{thm tm.pre_stren}, + post_weaken = @{thm tm.post_weaken}, + frame_rule = @{thm tm.frame_rule} + }:HoareTriple +*} + +ML {* + val _ = data_slot "HoareTriples" "HoareTriple list" "[]" +*} + +ML {* + val _ = HoareTriples_store [tm_triple] +*} + +ML {* (* aux1 functions *) + +fun focus_params t ctxt = + let + val (xs, Ts) = + split_list (Term.variant_frees t (Term.strip_all_vars t)); (*as they are printed :-*) + (* val (xs', ctxt') = variant_fixes xs ctxt; *) + (* val ps = xs' ~~ Ts; *) + val ps = xs ~~ Ts + val (_, ctxt'') = ctxt |> Variable.add_fixes xs + in ((xs, ps), ctxt'') end + +fun focus_concl ctxt t = + let + val ((xs, ps), ctxt') = focus_params t ctxt + val t' = Term.subst_bounds (rev (map Free ps), Term.strip_all_body t); + in (t' |> Logic.strip_imp_concl, ctxt') end + + fun get_concl ctxt (i, state) = + nth (Thm.prems_of state) (i - 1) + |> focus_concl ctxt |> (fn (x, _) => x |> HOLogic.dest_Trueprop) + (* aux1 end *) +*} + +ML {* + fun indexing xs = upto (0, length xs - 1) ~~ xs + fun select_idxs idxs ps = + map_index (fn (i, e) => if (member (op =) idxs i) then [e] else []) ps |> flat + fun select_out_idxs idxs ps = + map_index (fn (i, e) => if (member (op =) idxs i) then [] else [e]) ps |> flat + fun match_pres ctxt mf env ps qs = + let fun sel_match mf env [] qs = [(env, [])] + | sel_match mf env (p::ps) qs = + let val pm = map (fn (i, q) => [(i, + let val _ = tracing ctxt "Matching:" + val _ = [p, q] |> + (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt + val r = mf p q env + in r end)] + handle _ => ( + let val _ = tracing ctxt "Failed matching:" + val _ = [p, q] |> + (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt + in [] end)) qs |> flat + val r = pm |> map (fn (i, env') => + let val qs' = filter_out (fn (j, q) => j = i) qs + in sel_match mf env' ps qs' |> + map (fn (env'', idxs) => (env'', i::idxs)) end) + |> flat + in r end + in sel_match mf env ps (indexing qs) end + + fun provable tac ctxt goal = + let + val _ = tracing ctxt "Provable trying to prove:" + val _ = [goal] |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt + in + (Goal.prove ctxt [] [] goal (fn {context, ...} => tac context 1); true) + handle exn => false + end + fun make_sense tac ctxt thm_assms env = + thm_assms |> map (inst env) |> forall (provable tac ctxt) +*} + +ML {* + fun triple_for ctxt goal = + filter (fn trpl => (#can_process trpl) ctxt goal) (HoareTriples.get (Proof_Context.theory_of ctxt)) |> hd + + fun step_terms_for thm goal ctxt = + let + val _ = tracing ctxt "This is the new version of step_terms_for!" + val _ = tracing ctxt "Tring to find triple processor: TP" + val TP = triple_for ctxt goal + val _ = #binding TP |> Binding.name_of |> tracing ctxt + fun mk_sep_conj tms = foldr (fn tm => fn rtm => + ((#sconj_term TP)$tm$rtm)) (but_last tms) (last tms) + val thm_concl = thm |> prop_of + |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop + val thm_assms = thm |> prop_of + |> Logic.strip_imp_prems + val cmd_pat = thm_concl |> #get_mid TP ctxt |> get_cmd ctxt + val cmd = goal |> #get_mid TP ctxt |> get_cmd ctxt + val _ = tracing ctxt "matching command ... " + val _ = tracing ctxt "cmd_pat = " + val _ = pterm ctxt cmd_pat + val (hcmd, env1, is_last) = (cmd, match ctxt cmd_pat cmd, true) + handle exn => (cmd |> get_hcmd, match ctxt cmd_pat (cmd |> get_hcmd), false) + val _ = tracing ctxt "hcmd =" + val _ = pterm ctxt hcmd + val _ = tracing ctxt "match command succeed! " + val _ = tracing ctxt "pres =" + val pres = goal |> #get_pre TP ctxt |> break_sep_conj + val _ = pres |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt + val _ = tracing ctxt "pre_pats =" + val pre_pats = thm_concl |> #get_pre TP ctxt |> inst env1 |> break_sep_conj + val _ = pre_pats |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt + val _ = tracing ctxt "post_pats =" + val post_pats = thm_concl |> #get_post TP ctxt |> inst env1 |> break_sep_conj + val _ = post_pats |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt + val _ = tracing ctxt "Calculating sols" + val sols = match_pres ctxt (match_env ctxt) env1 pre_pats pres + val _ = tracing ctxt "End calculating sols, sols =" + val _ = tracing ctxt (@{make_string} sols) + val _ = tracing ctxt "Calulating env2 and idxs" + val (env2, idxs) = filter (fn (env, idxs) => make_sense (#sep_conj_ac_tac TP) + ctxt thm_assms env) sols |> hd + val _ = tracing ctxt "End calculating env2 and idxs" + val _ = tracing ctxt "mterms =" + val mterms = select_idxs idxs pres |> map (inst env2) + val _ = mterms |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt + val _ = tracing ctxt "nmterms = " + val nmterms = select_out_idxs idxs pres |> map (inst env2) + val _ = nmterms |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt + val pre_cond = pre_pats |> map (inst env2) |> mk_sep_conj + val post_cond = post_pats |> map (inst env2) |> mk_sep_conj + val post_cond_npc = + post_cond |> break_sep_conj |> filter (not_pred (#is_pc_term TP)) + |> (fn x => x @ nmterms) |> mk_sep_conj |> cterm_of (Proof_Context.theory_of ctxt) + fun mk_frame cond rest = + if rest = [] then cond else ((#sconj_term TP)$ cond) $ (mk_sep_conj rest) + val pre_cond_frame = mk_frame pre_cond nmterms |> cterm_of (Proof_Context.theory_of ctxt) + fun post_cond_frame j' = post_cond |> break_sep_conj |> filter (not_pred (#is_pc_term TP)) + |> (fn x => [#mk_pc_term TP j']@x) |> mk_sep_conj + |> (fn x => mk_frame x nmterms) + |> cterm_of (Proof_Context.theory_of ctxt) + val need_frame = (nmterms <> []) + in + (post_cond_npc, + pre_cond_frame, + post_cond_frame, need_frame, is_last) + end +*} + +ML {* + fun step_tac ctxt thm i state = + let + val _ = tracing ctxt "This is the new version of step_tac" + val (goal, ctxt) = nth (Thm.prems_of state) (i - 1) + |> focus_concl ctxt + |> (apfst HOLogic.dest_Trueprop) + val _ = tracing ctxt "step_tac: goal = " + val _ = goal |> pterm ctxt + val _ = tracing ctxt "Start to calculate intermediate terms ... " + val (post_cond_npc, pre_cond_frame, post_cond_frame, need_frame, is_last) + = step_terms_for thm goal ctxt + val _ = tracing ctxt "Tring to find triple processor: TP" + val TP = triple_for ctxt goal + val _ = #binding TP |> Binding.name_of |> tracing ctxt + fun mk_sep_conj tms = foldr (fn tm => fn rtm => + ((#sconj_term TP)$tm$rtm)) (but_last tms) (last tms) + val _ = tracing ctxt "Calculate intermediate terms finished! " + val post_cond_npc_str = post_cond_npc |> string_for_cterm ctxt + val pre_cond_frame_str = pre_cond_frame |> string_for_cterm ctxt + val _ = tracing ctxt "step_tac: post_cond_npc = " + val _ = post_cond_npc |> pcterm ctxt + val _ = tracing ctxt "step_tac: pre_cond_frame = " + val _ = pre_cond_frame |> pcterm ctxt + fun tac1 i state = + if is_last then (K all_tac) i state else + res_inst_tac ctxt [(("q", 0), post_cond_npc_str)] + (#hoare_seq1 TP) i state + fun tac2 i state = res_inst_tac ctxt [(("p", 0), pre_cond_frame_str)] + (#pre_stren TP) i state + fun foc_tac post_cond_frame ctxt i state = + let + val goal = get_concl ctxt (i, state) + val pc_term = goal |> #get_post TP ctxt |> break_sep_conj + |> filter (#is_pc_term TP) |> hd + val (_$Free(j', _)) = pc_term + val psd = post_cond_frame j' + val str_psd = psd |> string_for_cterm ctxt + val _ = tracing ctxt "foc_tac: psd = " + val _ = psd |> pcterm ctxt + in + res_inst_tac ctxt [(("q", 0), str_psd)] + (#post_weaken TP) i state + end + val frame_tac = if need_frame then (rtac (#frame_rule TP)) else (K all_tac) + val print_tac = if (Config.get ctxt trace_step) then Tactical.print_tac else (K all_tac) + val tac = (tac1 THEN' (K (print_tac "tac1 success"))) THEN' + (tac2 THEN' (K (print_tac "tac2 success"))) THEN' + ((foc_tac post_cond_frame ctxt) THEN' (K (print_tac "foc_tac success"))) THEN' + (frame_tac THEN' (K (print_tac "frame_tac success"))) THEN' + (((rtac thm) THEN_ALL_NEW (#sep_conj_ac_tac TP ctxt)) THEN' (K (print_tac "rtac thm success"))) THEN' + (K (ALLGOALS (atomic (#sep_conj_ac_tac TP ctxt)))) THEN' + (* (#sep_conj_ac_tac TP ctxt) THEN' (#sep_conj_ac_tac TP ctxt) THEN' *) + (K prune_params_tac) + in + tac i state + end + + fun unfold_cell_tac ctxt = (Local_Defs.unfold_tac ctxt @{thms one_def zero_def}) + fun fold_cell_tac ctxt = (Local_Defs.fold_tac ctxt @{thms one_def zero_def}) +*} + +ML {* + fun sg_step_tac thms ctxt = + let val sg_step_tac' = (map (fn thm => attemp (step_tac ctxt thm)) thms) + (* @ [attemp (goto_tac ctxt)] *) + |> FIRST' + val sg_step_tac'' = (K (unfold_cell_tac ctxt)) THEN' sg_step_tac' THEN' (K (fold_cell_tac ctxt)) + in + sg_step_tac' ORELSE' sg_step_tac'' + end + fun steps_tac thms ctxt i = REPEAT (sg_step_tac thms ctxt i) THEN (prune_params_tac) +*} + +ML {* + open StackMonad +*} + +method_setup hstep = {* + Attrib.thms >> (fn thms => fn ctxt => + (SIMPLE_METHOD' (fn i => + sg_step_tac (thms@(StepRules.get ctxt)) ctxt i))) + *} + "One step symbolic execution using step theorems." + +method_setup hsteps = {* + Attrib.thms >> (fn thms => fn ctxt => + (SIMPLE_METHOD' (fn i => + steps_tac (thms@(StepRules.get ctxt)) ctxt i))) + *} + "Sequential symbolic execution using step theorems." + +ML {* + fun goto_tac ctxt thm i state = + let + val (goal, ctxt) = nth (Thm.prems_of state) (i - 1) + |> focus_concl ctxt |> (apfst HOLogic.dest_Trueprop) + val _ = tracing ctxt "goto_tac: goal = " + val _ = goal |> string_of_term ctxt |> tracing ctxt + val (post_cond_npc, pre_cond_frame, post_cond_frame, need_frame, is_last) + = step_terms_for thm goal ctxt + val _ = tracing ctxt "Tring to find triple processor: TP" + val TP = triple_for ctxt goal + val _ = #binding TP |> Binding.name_of |> tracing ctxt + val _ = tracing ctxt "aaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaa" + val post_cond_npc_str = post_cond_npc |> string_for_cterm ctxt + val pre_cond_frame_str = pre_cond_frame |> string_for_cterm ctxt + val _ = tracing ctxt "goto_tac: post_cond_npc = " + val _ = post_cond_npc_str |> tracing ctxt + val _ = tracing ctxt "goto_tac: pre_cond_frame = " + val _ = pre_cond_frame_str |> tracing ctxt + fun tac1 i state = + if is_last then (K all_tac) i state else + res_inst_tac ctxt [] + (#hoare_seq2 TP) i state + fun tac2 i state = res_inst_tac ctxt [(("p", 0), pre_cond_frame_str)] + (#pre_stren TP) i state + fun foc_tac post_cond_frame ctxt i state = + let + val goal = get_concl ctxt (i, state) + val pc_term = goal |> #get_post TP ctxt |> break_sep_conj + |> filter (#is_pc_term TP) |> hd + val (_$Free(j', _)) = pc_term + val psd = post_cond_frame j' + val str_psd = psd |> string_for_cterm ctxt + val _ = tracing ctxt "goto_tac: psd = " + val _ = str_psd |> tracing ctxt + in + res_inst_tac ctxt [(("q", 0), str_psd)] + (#post_weaken TP) i state + end + val frame_tac = if need_frame then (rtac (#frame_rule TP)) else (K all_tac) + val _ = tracing ctxt "goto_tac: starting to apply tacs" + val print_tac = if (Config.get ctxt trace_step) then Tactical.print_tac else (K all_tac) + val tac = (tac1 THEN' (K (print_tac "tac1 success"))) THEN' + (tac2 THEN' (K (print_tac "tac2 success"))) THEN' + ((foc_tac post_cond_frame ctxt) THEN' (K (print_tac "foc_tac success"))) THEN' + (frame_tac THEN' (K (print_tac "frame_tac success"))) THEN' + ((((rtac thm) THEN_ALL_NEW (#sep_conj_ac_tac TP ctxt))) THEN' + (K (print_tac "rtac success")) + ) THEN' + (K (ALLGOALS (atomic (#sep_conj_ac_tac TP ctxt)))) THEN' + (K prune_params_tac) + in + tac i state + end +*} + +ML {* + fun sg_goto_tac thms ctxt = + let val sg_goto_tac' = (map (fn thm => attemp (goto_tac ctxt thm)) thms) + |> FIRST' + val sg_goto_tac'' = (K (unfold_cell_tac ctxt)) THEN' sg_goto_tac' THEN' (K (fold_cell_tac ctxt)) + in + sg_goto_tac' ORELSE' sg_goto_tac'' + end + fun gotos_tac thms ctxt i = REPEAT (sg_goto_tac thms ctxt i) THEN (prune_params_tac) +*} + +method_setup hgoto = {* + Attrib.thms >> (fn thms => fn ctxt => + (SIMPLE_METHOD' (fn i => + sg_goto_tac (thms@(StepRules.get ctxt)) ctxt i))) + *} + "One step symbolic execution using goto theorems." + +subsection {* Tactic for forward reasoning *} + +ML {* +fun mk_msel_rule ctxt conclusion idx term = +let + val cjt_count = term |> break_sep_conj |> length + fun variants nctxt names = fold_map Name.variant names nctxt; + + val (state, nctxt0) = Name.variant "s" (Variable.names_of ctxt); + + fun sep_conj_prop cjts = + FunApp.fun_app_free + (FunApp.fun_app_foldr SepConj.sep_conj_term cjts) state + |> HOLogic.mk_Trueprop; + + (* concatenate string and string of an int *) + fun conc_str_int str int = str ^ Int.toString int; + + (* make the conjunct names *) + val (cjts, _) = ListExtra.range 1 cjt_count + |> map (conc_str_int "a") |> variants nctxt0; + + fun skel_sep_conj names (Const (@{const_name sep_conj}, _) $ t1 $ t2 $ y) = + (let val nm1 = take (length (break_sep_conj t1)) names + val nm2 = drop (length (break_sep_conj t1)) names + val t1' = skel_sep_conj nm1 t1 + val t2' = skel_sep_conj nm2 t2 + in (SepConj.sep_conj_term $ t1' $ t2' $ y) end) + | skel_sep_conj names (Const (@{const_name sep_conj}, _) $ t1 $ t2) = + (let val nm1 = take (length (break_sep_conj t1)) names + val nm2 = drop (length (break_sep_conj t1)) names + val t1' = skel_sep_conj nm1 t1 + val t2' = skel_sep_conj nm2 t2 + in (SepConj.sep_conj_term $ t1' $ t2') end) + | skel_sep_conj names (Abs (x, y, t $ Bound 0)) = + let val t' = (skel_sep_conj names t) + val ty' = t' |> type_of |> domain_type + in (Abs (x, ty', (t' $ Bound 0))) end + | skel_sep_conj names t = Free (hd names, SepConj.sep_conj_term |> type_of |> domain_type); + val _ = tracing ctxt "xxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxx" + val oskel = skel_sep_conj cjts term; + val _ = tracing ctxt "yyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyy" + val ttt = oskel |> type_of + val _ = tracing ctxt "zzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzz" + val orig = FunApp.fun_app_free oskel state |> HOLogic.mk_Trueprop + val _ = tracing ctxt "uuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuu" + val is_selected = member (fn (x, y) => x = y) idx + val all_idx = ListExtra.range 0 cjt_count + val selected_idx = idx + val unselected_idx = filter_out is_selected all_idx + val selected = map (nth cjts) selected_idx + val unselected = map (nth cjts) unselected_idx + + fun fun_app_foldr f [a,b] = FunApp.fun_app_free (FunApp.fun_app_free f a) b + | fun_app_foldr f [a] = Free (a, SepConj.sep_conj_term |> type_of |> domain_type) + | fun_app_foldr f (x::xs) = (FunApp.fun_app_free f x) $ (fun_app_foldr f xs) + | fun_app_foldr _ _ = raise Fail "fun_app_foldr"; + + val reordered_skel = + if unselected = [] then (fun_app_foldr SepConj.sep_conj_term selected) + else (SepConj.sep_conj_term $ (fun_app_foldr SepConj.sep_conj_term selected) + $ (fun_app_foldr SepConj.sep_conj_term unselected)) + + val reordered = FunApp.fun_app_free reordered_skel state |> HOLogic.mk_Trueprop + val goal = Logic.mk_implies + (if conclusion then (orig, reordered) else (reordered, orig)); + val rule = + Goal.prove ctxt [] [] goal (fn _ => + auto_tac (ctxt |> Simplifier.map_simpset (fn ss => ss addsimps @{thms sep_conj_ac}))) + |> Drule.export_without_context +in + rule +end +*} + +lemma fwd_rule: + assumes "\ s . U s \ V s" + shows "(U ** RR) s \ (V ** RR) s" + by (metis assms sep_globalise) + +ML {* + fun sg_sg_fwd_tac ctxt thm pos i state = + let + + val tracing = (fn str => + if (Config.get ctxt trace_fwd) then Output.tracing str else ()) + fun pterm t = + t |> string_of_term ctxt |> tracing + fun pcterm ct = ct |> string_of_cterm ctxt |> tracing + + fun atm thm = + let + (* val thm = thm |> Drule.forall_intr_vars *) + val res = thm |> cprop_of |> Object_Logic.atomize + val res' = Raw_Simplifier.rewrite_rule [res] thm + in res' end + + fun find_idx ctxt pats terms = + let val result = + map (fn pat => (find_index (fn trm => ((match ctxt pat trm; true) + handle _ => false)) terms)) pats + in (assert_all (fn x => x >= 0) result (K "match of precondition failed")); + result + end + + val goal = nth (Drule.cprems_of state) (i - 1) |> term_of + val _ = tracing "goal = " + val _ = goal |> pterm + + val ctxt_orig = ctxt + + val ((ps, goal), ctxt) = Variable.focus goal ctxt_orig + + val prems = goal |> Logic.strip_imp_prems + + val cprem = nth prems (pos - 1) + val (_ $ (the_prem $ _)) = cprem + val cjts = the_prem |> break_sep_conj + val thm_prems = thm |> cprems_of |> hd |> Thm.dest_arg |> Thm.dest_fun + val thm_assms = thm |> cprems_of |> tl |> map term_of + val thm_cjts = thm_prems |> term_of |> break_sep_conj + val thm_trm = thm |> prop_of + + val _ = tracing "cjts = " + val _ = cjts |> map pterm + val _ = tracing "thm_cjts = " + val _ = thm_cjts |> map pterm + + val _ = tracing "Calculating sols" + val sols = match_pres ctxt (match_env ctxt) empty_env thm_cjts cjts + val _ = tracing "End calculating sols, sols =" + val _ = tracing (@{make_string} sols) + val _ = tracing "Calulating env2 and idxs" + val (env2, idx) = filter (fn (env, idxs) => make_sense sep_conj_ac_tac ctxt thm_assms env) sols |> hd + val ([thm'_trm], ctxt') = thm_trm |> inst env2 |> single + |> (fn trms => Variable.import_terms true trms ctxt) + val thm'_prem = Logic.strip_imp_prems thm'_trm |> hd + val thm'_concl = Logic.strip_imp_concl thm'_trm + val thm'_prem = (Goal.prove ctxt' [] [thm'_prem] thm'_concl + (fn {context, prems = [prem]} => + (rtac (prem RS thm) THEN_ALL_NEW (sep_conj_ac_tac ctxt)) 1)) + val [thm'] = Variable.export ctxt' ctxt_orig [thm'_prem] + val trans_rule = + mk_msel_rule ctxt true idx the_prem + val _ = tracing "trans_rule = " + val _ = trans_rule |> cprop_of |> pcterm + val app_rule = + if (length cjts = length thm_cjts) then thm' else + ((thm' |> atm) RS @{thm fwd_rule}) + val _ = tracing "app_rule = " + val _ = app_rule |> cprop_of |> pcterm + val print_tac = if (Config.get ctxt trace_fwd) then Tactical.print_tac else (K all_tac) + val the_tac = (dtac trans_rule THEN' (K (print_tac "dtac1 success"))) THEN' + ((dtac app_rule THEN' (K (print_tac "dtac2 success")))) +in + (the_tac i state) handle _ => no_tac state +end +*} + +ML {* + fun sg_fwd_tac ctxt thm i state = + let + val goal = nth (Drule.cprems_of state) (i - 1) + val prems = goal |> term_of |> Term.strip_all_body |> Logic.strip_imp_prems + val posx = ListExtra.range 1 (length prems) + in + ((map (fn pos => attemp (sg_sg_fwd_tac ctxt thm pos)) posx) |> FIRST') i state + end + + fun fwd_tac ctxt thms i state = + ((map (fn thm => sg_fwd_tac ctxt thm) thms) |> FIRST') i state +*} + +method_setup fwd = {* + Attrib.thms >> (fn thms => fn ctxt => + (SIMPLE_METHOD' (fn i => + fwd_tac ctxt (thms@(FwdRules.get ctxt)) i))) + *} + "Forward derivation of separation implication" + +text {* Testing the fwd tactic *} + +lemma ones_abs: + assumes "(ones u v \* ones w x) s" "w = v + 1" + shows "ones u x s" + using assms(1) unfolding assms(2) +proof(induct u v arbitrary: x s rule:ones_induct) + case (Base i j x s) + thus ?case by (auto elim!:condE) +next + case (Step i j x s) + hence h: "\ x s. (ones (i + 1) j \* ones (j + 1) x) s \ ones (i + 1) x s" + by metis + hence "(ones (i + 1) x \* one i) s" + by (rule fwd_rule, insert Step(3), auto simp:sep_conj_ac) + thus ?case + by (smt condD ones.simps sep_conj_commute) +qed + +lemma one_abs: "(one m) s \ (ones m m) s" + by (smt cond_true_eq2 ones.simps) + +lemma ones_reps_abs: + assumes "ones m n s" + "m \ n" + shows "(reps m n [nat (n - m)]) s" + using assms + by simp + +lemma reps_reps'_abs: + assumes "(reps m n xs \* zero u) s" "u = n + 1" "xs \ []" + shows "(reps' m u xs) s" + unfolding assms using assms + by (unfold reps'_def, simp) + +lemma reps'_abs: + assumes "(reps' m n xs \* reps' u v ys) s" "u = n + 1" + shows "(reps' m v (xs @ ys)) s" + apply (unfold reps'_append, rule_tac x = u in EXS_intro) + by (insert assms, simp) + +lemmas abs_ones = one_abs ones_abs + +lemmas abs_reps' = ones_reps_abs reps_reps'_abs reps'_abs + + +section {* Modular TM programming and verification *} + +lemma ones_false [simp]: "j < i - 1 \ (ones i j) = sep_false" + by (simp add:pasrt_def) + +lemma hoare_right_until_zero: + "\st i ** ps u ** ones u (v - 1) ** zero v \ + i:[right_until_zero]:j + \st j ** ps v ** ones u (v - 1) ** zero v \" +proof(unfold right_until_zero_def, + intro t_hoare_local t_hoare_label, clarify, + rule t_hoare_label_last, simp, simp) + fix la + let ?body = "i :[ (if_zero la ; move_right ; jmp i) ]: la" + let ?j = la + show "\st i \* ps u \* ones u (v - 1) \* zero v\ ?body + \st ?j \* ps v \* ones u (v - 1) \* zero v\" (is "?P u (v - 1) (ones u (v - 1))") + proof(induct "u" "v - 1" rule:ones_induct) + case (Base k) + moreover have "\st i \* ps v \* zero v\ ?body + \st ?j \* ps v \* zero v\" by hsteps + ultimately show ?case by (auto intro!:tm.pre_condI simp:sep_conj_cond) + next + case (Step k) + moreover have "\st i \* ps k \* (one k \* ones (k + 1) (v - 1)) \* zero v\ + i :[ (if_zero ?j ; move_right ; jmp i) ]: ?j + \st ?j \* ps v \* (one k \* ones (k + 1) (v - 1)) \* zero v\" + proof - + have s1: "\st i \* ps k \* (one k \* ones (k + 1) (v - 1)) \* zero v\ + ?body + \st i \* ps (k + 1) \* one k \* ones (k + 1) (v - 1) \* zero v\" + proof(cases "k + 1 \ v") + case True + with Step(1) have "v = k + 1" by arith + thus ?thesis + apply(simp add: one_def) + by hsteps + next + case False + hence eq_ones: "ones (k + 1) (v - 1) = + (one (k + 1) \* ones ((k + 1) + 1) (v - 1))" + by simp + show ?thesis + apply(simp only: eq_ones) + by hsteps + qed + note Step(2)[step] + have s2: "\st i \* ps (k + 1) \* one k \* ones (k + 1) (v - 1) \* zero v\ + ?body + \st ?j \* ps v \* one k \* ones (k + 1) (v - 1) \* zero v\" + by hsteps + from tm.sequencing [OF s1 s2, step] + show ?thesis + by (auto simp:sep_conj_ac) + qed + ultimately show ?case by simp + qed +qed + +lemma hoare_right_until_zero_gen[step]: + assumes "u = v" "w = x - 1" + shows "\st i ** ps u ** ones v w ** zero x \ + i:[right_until_zero]:j + \st j ** ps x ** ones v w ** zero x \" + by (unfold assms, rule hoare_right_until_zero) + +lemma hoare_left_until_zero: + "\st i ** ps v ** zero u ** ones (u + 1) v \ + i:[left_until_zero]:j + \st j ** ps u ** zero u ** ones (u + 1) v \" +proof(unfold left_until_zero_def, + intro t_hoare_local t_hoare_label, clarify, + rule t_hoare_label_last, simp+) + fix la + let ?body = "i :[ (if_zero la ; move_left ; jmp i) ]: la" + let ?j = la + show "\st i \* ps v \* zero u \* ones (u + 1) v\ ?body + \st ?j \* ps u \* zero u \* ones (u + 1) v\" + proof(induct "u+1" v rule:ones_rev_induct) + case (Base k) + thus ?case + by (simp add:sep_conj_cond, intro tm.pre_condI, simp, hstep) + next + case (Step k) + have "\st i \* ps k \* zero u \* ones (u + 1) (k - 1) \* one k\ + ?body + \st ?j \* ps u \* zero u \* ones (u + 1) (k - 1) \* one k\" + proof(rule tm.sequencing[where q = + "st i \* ps (k - 1) \* zero u \* ones (u + 1) (k - 1) \* one k"]) + show "\st i \* ps k \* zero u \* ones (u + 1) (k - 1) \* one k\ + ?body + \st i \* ps (k - 1) \* zero u \* ones (u + 1) (k - 1) \* one k\" + proof(induct "u + 1" "k - 1" rule:ones_rev_induct) + case Base with Step(1) have "k = u + 1" by arith + thus ?thesis + by (simp, hsteps) + next + case Step + show ?thesis + apply (unfold ones_rev[OF Step(1)], simp) + apply (unfold one_def) + by hsteps + qed + next + note Step(2) [step] + show "\st i \* ps (k - 1) \* zero u \* ones (u + 1) (k - 1) \* one k\ + ?body + \st ?j \* ps u \* zero u \* ones (u + 1) (k - 1) \* one k\" by hsteps + qed + thus ?case by (unfold ones_rev[OF Step(1)], simp) + qed +qed + +lemma hoare_left_until_zero_gen[step]: + assumes "u = x" "w = v + 1" + shows "\st i ** ps u ** zero v ** ones w x \ + i:[left_until_zero]:j + \st j ** ps v ** zero v ** ones w x \" + by (unfold assms, rule hoare_left_until_zero) + +lemma hoare_right_until_one: + "\st i ** ps u ** zeros u (v - 1) ** one v \ + i:[right_until_one]:j + \st j ** ps v ** zeros u (v - 1) ** one v \" +proof(unfold right_until_one_def, + intro t_hoare_local t_hoare_label, clarify, + rule t_hoare_label_last, simp+) + fix la + let ?body = "i :[ (if_one la ; move_right ; jmp i) ]: la" + let ?j = la + show "\st i \* ps u \* zeros u (v - 1) \* one v\ ?body + \st ?j \* ps v \* zeros u (v - 1) \* one v\" + proof(induct u "v - 1" rule:zeros_induct) + case (Base k) + thus ?case + by (simp add:sep_conj_cond, intro tm.pre_condI, simp, hsteps) + next + case (Step k) + have "\st i \* ps k \* zero k \* zeros (k + 1) (v - 1) \* one v\ + ?body + \st ?j \* ps v \* zero k \* zeros (k + 1) (v - 1) \* one v\" + proof(rule tm.sequencing[where q = + "st i \* ps (k + 1) \* zero k \* zeros (k + 1) (v - 1) \* one v"]) + show "\st i \* ps k \* zero k \* zeros (k + 1) (v - 1) \* one v\ + ?body + \st i \* ps (k + 1) \* zero k \* zeros (k + 1) (v - 1) \* one v\" + proof(induct "k + 1" "v - 1" rule:zeros_induct) + case Base + with Step(1) have eq_v: "k + 1 = v" by arith + from Base show ?thesis + apply (simp add:sep_conj_cond, intro tm.pre_condI, simp) + apply (hstep, clarsimp) + by hsteps + next + case Step + thus ?thesis + by (simp, hsteps) + qed + next + note Step(2)[step] + show "\st i \* ps (k + 1) \* zero k \* zeros (k + 1) (v - 1) \* one v\ + ?body + \st ?j \* ps v \* zero k \* zeros (k + 1) (v - 1) \* one v\" + by hsteps + qed + thus ?case by (auto simp: sep_conj_ac Step(1)) + qed +qed + +lemma hoare_right_until_one_gen[step]: + assumes "u = v" "w = x - 1" + shows + "\st i ** ps u ** zeros v w ** one x \ + i:[right_until_one]:j + \st j ** ps x ** zeros v w ** one x \" + by (unfold assms, rule hoare_right_until_one) + +lemma hoare_left_until_one: + "\st i ** ps v ** one u ** zeros (u + 1) v \ + i:[left_until_one]:j + \st j ** ps u ** one u ** zeros (u + 1) v \" +proof(unfold left_until_one_def, + intro t_hoare_local t_hoare_label, clarify, + rule t_hoare_label_last, simp+) + fix la + let ?body = "i :[ (if_one la ; move_left ; jmp i) ]: la" + let ?j = la + show "\st i \* ps v \* one u \* zeros (u + 1) v\ ?body + \st ?j \* ps u \* one u \* zeros (u + 1) v\" + proof(induct u v rule: ones'.induct) + fix ia ja + assume h: "\ ja < ia \ + \st i \* ps (ja - 1) \* one ia \* zeros (ia + 1) (ja - 1)\ ?body + \st ?j \* ps ia \* one ia \* zeros (ia + 1) (ja - 1)\" + show "\st i \* ps ja \* one ia \* zeros (ia + 1) ja\ ?body + \st ?j \* ps ia \* one ia \* zeros (ia + 1) ja\" + proof(cases "ja < ia") + case False + note lt = False + from h[OF this] have [step]: + "\st i \* ps (ja - 1) \* one ia \* zeros (ia + 1) (ja - 1)\ ?body + \st ?j \* ps ia \* one ia \* zeros (ia + 1) (ja - 1)\" . + show ?thesis + proof(cases "ja = ia") + case True + moreover + have "\st i \* ps ja \* one ja\ ?body \st ?j \* ps ja \* one ja\" + by hsteps + ultimately show ?thesis by auto + next + case False + with lt have k1: "ia < ja" by auto + from zeros_rev[of "ja" "ia + 1"] this + have eq_zeros: "zeros (ia + 1) ja = (zeros (ia + 1) (ja - 1) \* zero ja)" + by simp + have s1: "\st i \* ps ja \* one ia \* zeros (ia + 1) (ja - 1) \* zero ja\ + ?body + \st i \* ps (ja - 1) \* one ia \* zeros (ia + 1) (ja - 1) \* zero ja\" + proof(cases "ia + 1 \ ja") + case True + from k1 True have "ja = ia + 1" by arith + moreover have "\st i \* ps (ia + 1) \* one (ia + 1 - 1) \* zero (ia + 1)\ + i :[ (if_one ?j ; move_left ; jmp i) ]: ?j + \st i \* ps (ia + 1 - 1) \* one (ia + 1 - 1) \* zero (ia + 1)\" + by (hsteps) + ultimately show ?thesis + by (simp) + next + case False + from zeros_rev[of "ja - 1" "ia + 1"] False + have k: "zeros (ia + 1) (ja - 1) = + (zeros (ia + 1) (ja - 1 - 1) \* zero (ja - 1))" + by auto + show ?thesis + apply (unfold k, simp) + by hsteps + qed + have s2: "\st i \* ps (ja - 1) \* one ia \* zeros (ia + 1) (ja - 1) \* zero ja\ + ?body + \st ?j \* ps ia \* one ia \* zeros (ia + 1) (ja - 1) \* zero ja\" + by hsteps + from tm.sequencing[OF s1 s2, step] + show ?thesis + apply (unfold eq_zeros) + by hstep + qed (* ccc *) + next + case True + thus ?thesis by (auto intro:tm.hoare_sep_false) + qed + qed +qed + +lemma hoare_left_until_one_gen[step]: + assumes "u = x" "w = v + 1" + shows "\st i ** ps u ** one v ** zeros w x \ + i:[left_until_one]:j + \st j ** ps v ** one v ** zeros w x \" + by (unfold assms, rule hoare_left_until_one) + +declare ones.simps[simp del] + +lemma reps_simps3: "ks \ [] \ + reps i j (k # ks) = (ones i (i + int k) ** zero (i + int k + 1) ** reps (i + int k + 2) j ks)" +by(case_tac ks, simp, simp add: reps.simps) + +lemma cond_eqI: + assumes h: "b \ r = s" + shows "( ** r) = ( ** s)" +proof(cases b) + case True + from h[OF this] show ?thesis by simp +next + case False + thus ?thesis + by (unfold sep_conj_def set_ins_def pasrt_def, auto) +qed + +lemma reps_rev: "ks \ [] + \ reps i j (ks @ [k]) = (reps i (j - int (k + 1) - 1 ) ks \* + zero (j - int (k + 1)) \* ones (j - int k) j)" +proof(induct ks arbitrary: i j) + case Nil + thus ?case by simp +next + case (Cons a ks) + show ?case + proof(cases "ks = []") + case True + thus ?thesis + proof - + have eq_cond: "(j = i + int a + 2 + int k) = (-2 + (j - int k) = i + int a)" by auto + have "(<(-2 + (j - int k) = i + int a)> \* + one i \* ones (i + 1) (i + int a) \* + zero (i + int a + 1) \* one (i + int a + 2) \* ones (3 + (i + int a)) (i + int a + 2 + int k)) + = + (<(-2 + (j - int k) = i + int a)> \* one i \* ones (i + 1) (i + int a) \* + zero (j - (1 + int k)) \* one (j - int k) \* ones (j - int k + 1) j)" + (is "( \* ?L) = ( \* ?R)") + proof(rule cond_eqI) + assume h: "-2 + (j - int k) = i + int a" + hence eqs: "i + int a + 1 = j - (1 + int k)" + "i + int a + 2 = j - int k" + "3 + (i + int a) = j - int k + 1" + "(i + int a + 2 + int k) = j" + by auto + show "?L = ?R" + by (unfold eqs, auto simp:sep_conj_ac) + qed + with True + show ?thesis + apply (simp del:ones_simps reps.simps) + apply (simp add:sep_conj_cond eq_cond) + by (auto simp:sep_conj_ac) + qed + next + case False + from Cons(1)[OF False, of "i + int a + 2" j] this + show ?thesis + by(simp add: reps_simps3 sep_conj_ac) + qed +qed + +lemma hoare_if_one_reps: + assumes nn: "ks \ []" + shows "\st i ** ps v ** reps u v ks\ + i:[if_one e]:j + \st e ** ps v ** reps u v ks\" +proof(rule rev_exhaust[of ks]) + assume "ks = []" with nn show ?thesis by simp +next + fix y ys + assume eq_ks: "ks = ys @ [y]" + show " \st i \* ps v \* reps u v ks\ i :[ if_one e ]: j \st e \* ps v \* reps u v ks\" + proof(cases "ys = []") + case False + have "\st i \* ps v \* reps u v (ys @ [y])\ i :[ if_one e ]: j \st e \* ps v \* reps u v (ys @ [y])\" + apply(unfold reps_rev[OF False], simp del:ones_simps add:ones_rev) + by hstep + thus ?thesis + by (simp add:eq_ks) + next + case True + with eq_ks + show ?thesis + apply (simp del:ones_simps add:ones_rev sep_conj_cond, intro tm.pre_condI, simp) + by hstep + qed +qed + +lemma hoare_if_one_reps_gen[step]: + assumes nn: "ks \ []" "u = w" + shows "\st i ** ps u ** reps v w ks\ + i:[if_one e]:j + \st e ** ps u ** reps v w ks\" + by (unfold `u = w`, rule hoare_if_one_reps[OF `ks \ []`]) + +lemma hoare_if_zero_ones_false[step]: + assumes "\ w < u" "v = w" + shows "\st i \* ps v \* ones u w\ + i :[if_zero e]: j + \st j \* ps v \* ones u w\" + by (unfold `v = w` ones_rev[OF `\ w < u`], hstep) + +lemma hoare_left_until_double_zero_nil[step]: + assumes "u = v" + shows "\st i ** ps u ** zero v\ + i:[left_until_double_zero]:j + \st j ** ps u ** zero v\" + apply (unfold `u = v` left_until_double_zero_def, + intro t_hoare_local t_hoare_label, clarsimp, + rule t_hoare_label_last, simp+) + by (hsteps) + +lemma hoare_if_zero_reps_false: + assumes nn: "ks \ []" + shows "\st i ** ps v ** reps u v ks\ + i:[if_zero e]:j + \st j ** ps v ** reps u v ks\" +proof(rule rev_exhaust[of ks]) + assume "ks = []" with nn show ?thesis by simp +next + fix y ys + assume eq_ks: "ks = ys @ [y]" + show " \st i \* ps v \* reps u v ks\ i :[ if_zero e ]: j \st j \* ps v \* reps u v ks\" + proof(cases "ys = []") + case False + have "\st i \* ps v \* reps u v (ys @ [y])\ i :[ if_zero e ]: j \st j \* ps v \* reps u v (ys @ [y])\" + apply(unfold reps_rev[OF False], simp del:ones_simps add:ones_rev) + by hstep + thus ?thesis + by (simp add:eq_ks) + next + case True + with eq_ks + show ?thesis + apply (simp del:ones_simps add:ones_rev sep_conj_cond, intro tm.pre_condI, simp) + by hstep + qed +qed + +lemma hoare_if_zero_reps_false_gen[step]: + assumes "ks \ []" "u = w" + shows "\st i ** ps u ** reps v w ks\ + i:[if_zero e]:j + \st j ** ps u ** reps v w ks\" + by (unfold `u = w`, rule hoare_if_zero_reps_false[OF `ks \ []`]) + + +lemma hoare_if_zero_reps_false1: + assumes nn: "ks \ []" + shows "\st i ** ps u ** reps u v ks\ + i:[if_zero e]:j + \st j ** ps u ** reps u v ks\" +proof - + from nn obtain y ys where eq_ys: "ks = y#ys" + by (metis neq_Nil_conv) + show ?thesis + apply (unfold eq_ys) + by (case_tac ys, (simp, hsteps)+) +qed + +lemma hoare_if_zero_reps_false1_gen[step]: + assumes nn: "ks \ []" + and h: "u = w" + shows "\st i ** ps u ** reps w v ks\ + i:[if_zero e]:j + \st j ** ps u ** reps w v ks\" + by (unfold h, rule hoare_if_zero_reps_false1[OF `ks \ []`]) + +lemma hoare_left_until_double_zero: + assumes h: "ks \ []" + shows "\st i ** ps v ** zero u ** zero (u + 1) ** reps (u+2) v ks\ + i:[left_until_double_zero]:j + \st j ** ps u ** zero u ** zero (u + 1) ** reps (u+2) v ks\" +proof(unfold left_until_double_zero_def, + intro t_hoare_local t_hoare_label, clarsimp, + rule t_hoare_label_last, simp+) + fix la + let ?body = "i :[ (if_zero la ; left_until_zero ; move_left ; if_one i) ]: j" + let ?j = j + show "\st i \* ps v \* zero u \* zero (u + 1) \* reps (u + 2) v ks\ + ?body + \st ?j \* ps u \* zero u \* zero (u + 1) \* reps (u + 2) v ks\" + using h + proof(induct ks arbitrary: v rule:rev_induct) + case Nil + with h show ?case by auto + next + case (snoc k ks) + show ?case + proof(cases "ks = []") + case True + have eq_ones: + "ones (u + 2) (u + 2 + int k) = (ones (u + 2) (u + 1 + int k) \* one (u + 2 + int k))" + by (smt ones_rev) + have eq_ones': "(one (u + 2) \* ones (3 + u) (u + 2 + int k)) = + (one (u + 2 + int k) \* ones (u + 2) (u + 1 + int k))" + by (smt eq_ones ones.simps sep.mult_commute) + thus ?thesis + apply (insert True, simp del:ones_simps add:sep_conj_cond) + apply (rule tm.pre_condI, simp del:ones_simps, unfold eq_ones) + apply hsteps + apply (rule_tac p = "st j' \* ps (u + 2 + int k) \* zero u \* + zero (u + 1) \* ones (u + 2) (u + 2 + int k)" + in tm.pre_stren) + by (hsteps) + next + case False + from False have spt: "splited (ks @ [k]) ks [k]" by (unfold splited_def, auto) + show ?thesis + apply (unfold reps_splited[OF spt], simp del:ones_simps add:sep_conj_cond) + apply (rule tm.pre_condI, simp del:ones_simps) + apply (rule_tac q = "st i \* + ps (1 + (u + int (reps_len ks))) \* + zero u \* + zero (u + 1) \* + reps (u + 2) (1 + (u + int (reps_len ks))) ks \* + zero (u + 2 + int (reps_len ks)) \* + ones (3 + (u + int (reps_len ks))) (3 + (u + int (reps_len ks)) + int k)" in + tm.sequencing) + apply hsteps[1] + by (hstep snoc(1)) + qed + qed +qed + +lemma hoare_left_until_double_zero_gen[step]: + assumes h1: "ks \ []" + and h: "u = y" "w = v + 1" "x = v + 2" + shows "\st i ** ps u ** zero v ** zero w ** reps x y ks\ + i:[left_until_double_zero]:j + \st j ** ps v ** zero v ** zero w ** reps x y ks\" + by (unfold h, rule hoare_left_until_double_zero[OF h1]) + +lemma hoare_jmp_reps1: + assumes "ks \ []" + shows "\ st i \* ps u \* reps u v ks\ + i:[jmp e]:j + \ st e \* ps u \* reps u v ks\" +proof - + from assms obtain k ks' where Cons:"ks = k#ks'" + by (metis neq_Nil_conv) + thus ?thesis + proof(cases "ks' = []") + case True with Cons + show ?thesis + apply(simp add:sep_conj_cond reps.simps, intro tm.pre_condI, simp add:ones_simps) + by (hgoto hoare_jmp_gen) + next + case False + show ?thesis + apply (unfold `ks = k#ks'` reps_simp3[OF False], simp add:ones_simps) + by (hgoto hoare_jmp[where p = u]) + qed +qed + +lemma hoare_jmp_reps1_gen[step]: + assumes "ks \ []" "u = v" + shows "\ st i \* ps u \* reps v w ks\ + i:[jmp e]:j + \ st e \* ps u \* reps v w ks\" + by (unfold assms, rule hoare_jmp_reps1[OF `ks \ []`]) + +lemma hoare_jmp_reps: + "\ st i \* ps u \* reps u v ks \* tm (v + 1) x \ + i:[(jmp e; c)]:j + \ st e \* ps u \* reps u v ks \* tm (v + 1) x \" +proof(cases "ks") + case Nil + thus ?thesis + by (simp add:sep_conj_cond, intro tm.pre_condI, simp, hsteps) +next + case (Cons k ks') + thus ?thesis + proof(cases "ks' = []") + case True with Cons + show ?thesis + apply(simp add:sep_conj_cond, intro tm.pre_condI, simp) + by (hgoto hoare_jmp[where p = u]) + next + case False + show ?thesis + apply (unfold `ks = k#ks'` reps_simp3[OF False], simp) + by (hgoto hoare_jmp[where p = u]) + qed +qed + +lemma hoare_shift_right_cons: + assumes h: "ks \ []" + shows "\st i \* ps u ** reps u v ks \* zero (v + 1) \* zero (v + 2) \ + i:[shift_right]:j + \st j ** ps (v + 2) ** zero u ** reps (u + 1) (v + 1) ks ** zero (v + 2) \" +proof(unfold shift_right_def, intro t_hoare_local t_hoare_label, clarify, + rule t_hoare_label_last, auto) + fix la + have eq_ones: "\ u k. (one (u + int k + 1) \* ones (u + 1) (u + int k)) = + (one (u + 1) \* ones (2 + u) (u + 1 + int k))" + by (smt cond_true_eq2 ones.simps ones_rev sep.mult_assoc sep.mult_commute + sep.mult_left_commute sep_conj_assoc sep_conj_commute + sep_conj_cond1 sep_conj_cond2 sep_conj_cond3 sep_conj_left_commute + sep_conj_trivial_strip2) + show "\st i \* ps u \* reps u v ks \* zero (v + 1) \* zero (v + 2)\ + i :[ (if_zero la ; + write_zero ; move_right ; right_until_zero ; write_one ; move_right ; jmp i) ]: la + \st la \* ps (v + 2) \* zero u \* reps (u + 1) (v + 1) ks \* zero (v + 2)\" + using h + proof(induct ks arbitrary:i u v) + case (Cons k ks) + thus ?case + proof(cases "ks = []") + let ?j = la + case True + let ?body = "i :[ (if_zero ?j ; + write_zero ; + move_right ; + right_until_zero ; + write_one ; move_right ; jmp i) ]: ?j" + have first_interation: + "\st i \* ps u \* one u \* ones (u + 1) (u + int k) \* zero (u + int k + 1) \* + zero (u + int k + 2)\ + ?body + \st i \* + ps (u + int k + 2) \* + one (u + int k + 1) \* ones (u + 1) (u + int k) \* zero u \* zero (u + int k + 2)\" + apply (hsteps) + by (simp add:sep_conj_ac, sep_cancel+, smt) + hence "\st i \* ps u \* one u \* ones (u + 1) (u + int k) \* zero (u + int k + 1) \* + zero (u + int k + 2)\ + ?body + \st ?j \* ps (u + int k + 2) \* zero u \* one (u + 1) \* + ones (2 + u) (u + 1 + int k) \* zero (u + int k + 2)\" + proof(rule tm.sequencing) + show "\st i \* + ps (u + int k + 2) \* + one (u + int k + 1) \* ones (u + 1) (u + int k) \* zero u \* zero (u + int k + 2)\ + ?body + \st ?j \* + ps (u + int k + 2) \* + zero u \* one (u + 1) \* ones (2 + u) (u + 1 + int k) \* zero (u + int k + 2)\" + apply (hgoto hoare_if_zero_true_gen) + by (simp add:sep_conj_ac eq_ones) + qed + with True + show ?thesis + by (simp, simp only:sep_conj_cond, intro tm.pre_condI, auto simp:sep_conj_ac) + next + case False + let ?j = la + let ?body = "i :[ (if_zero ?j ; + write_zero ; + move_right ; right_until_zero ; + write_one ; move_right ; jmp i) ]: ?j" + have eq_ones': + "(one (u + int k + 1) \* + ones (u + 1) (u + int k) \* + zero u \* reps (u + int k + 2) v ks \* zero (v + 1) \* zero (v + 2)) + = + (zero u \* + ones (u + 1) (u + int k) \* + one (u + int k + 1) \* reps (u + int k + 2) v ks \* zero (v + 1) \* zero (v + 2))" + by (simp add:eq_ones sep_conj_ac) + have "\st i \* ps u \* one u \* ones (u + 1) (u + int k) \* zero (u + int k + 1) \* + reps (u + int k + 2) v ks \* zero (v + 1) \* zero (v + 2)\ + ?body + \st i \* ps (u + int k + 2) \* zero u \* ones (u + 1) (u + int k) \* + one (u + int k + 1) \* reps (u + int k + 2) v ks \* zero (v + 1) \* zero (v + 2)\" + apply (hsteps) + by (auto simp:sep_conj_ac, sep_cancel+, smt) + hence "\st i \* ps u \* one u \* ones (u + 1) (u + int k) \* zero (u + int k + 1) \* + reps (u + int k + 2) v ks \* zero (v + 1) \* zero (v + 2)\ + ?body + \st ?j \* ps (v + 2) \* zero u \* one (u + 1) \* ones (2 + u) (u + 1 + int k) \* + zero (2 + (u + int k)) \* reps (3 + (u + int k)) (v + 1) ks \* zero (v + 2)\" + proof(rule tm.sequencing) + have eq_ones': + "\ u k. (one (u + int k + 1) \* ones (u + 1) (u + int k) \* zero (u + int k + 2)) = + (one (u + 1) \* zero (2 + (u + int k)) \* ones (2 + u) (u + 1 + int k))" + by (smt eq_ones sep.mult_assoc sep_conj_commute) + show "\st i \* ps (u + int k + 2) \* zero u \* + ones (u + 1) (u + int k) \* one (u + int k + 1) \* reps (u + int k + 2) v ks \* + zero (v + 1) \* zero (v + 2)\ + ?body + \st ?j \* ps (v + 2) \* zero u \* one (u + 1) \* ones (2 + u) (u + 1 + int k) \* + zero (2 + (u + int k)) \* reps (3 + (u + int k)) (v + 1) ks \* zero (v + 2)\" + apply (hsteps Cons.hyps) + by (simp add:sep_conj_ac eq_ones, sep_cancel+, smt) + qed + thus ?thesis + by (unfold reps_simp3[OF False], auto simp:sep_conj_ac) + qed + qed auto +qed + +lemma hoare_shift_right_cons_gen[step]: + assumes h: "ks \ []" + and h1: "u = v" "x = w + 1" "y = w + 2" + shows "\st i \* ps u ** reps v w ks \* zero x \* zero y \ + i:[shift_right]:j + \st j ** ps y ** zero v ** reps (v + 1) x ks ** zero y\" + by (unfold h1, rule hoare_shift_right_cons[OF h]) + +lemma shift_right_nil [step]: + assumes "u = v" + shows + "\ st i \* ps u \* zero v \ + i:[shift_right]:j + \ st j \* ps u \* zero v \" + by (unfold assms shift_right_def, intro t_hoare_local t_hoare_label, clarify, + rule t_hoare_label_last, simp+, hstep) + + +text {* + @{text "clear_until_zero"} is useful to implement @{text "drag"}. +*} + +lemma hoare_clear_until_zero[step]: + "\st i ** ps u ** ones u v ** zero (v + 1)\ + i :[clear_until_zero]: j + \st j ** ps (v + 1) ** zeros u v ** zero (v + 1)\ " +proof(unfold clear_until_zero_def, intro t_hoare_local, rule t_hoare_label, + rule t_hoare_label_last, simp+) + let ?body = "i :[ (if_zero j ; write_zero ; move_right ; jmp i) ]: j" + show "\st i \* ps u \* ones u v \* zero (v + 1)\ ?body + \st j \* ps (v + 1) \* zeros u v \* zero (v + 1)\" + proof(induct u v rule: zeros.induct) + fix ia ja + assume h: "\ ja < ia \ + \st i \* ps (ia + 1) \* ones (ia + 1) ja \* zero (ja + 1)\ ?body + \st j \* ps (ja + 1) \* zeros (ia + 1) ja \* zero (ja + 1)\" + show "\st i \* ps ia \* ones ia ja \* zero (ja + 1)\ ?body + \st j \* ps (ja + 1) \* zeros ia ja \* zero (ja + 1)\" + proof(cases "ja < ia") + case True + thus ?thesis + by (simp add: ones.simps zeros.simps sep_conj_ac, simp only:sep_conj_cond, + intro tm.pre_condI, simp, hsteps) + next + case False + note h[OF False, step] + from False have ones_eq: "ones ia ja = (one ia \* ones (ia + 1) ja)" + by(simp add: ones.simps) + from False have zeros_eq: "zeros ia ja = (zero ia \* zeros (ia + 1) ja)" + by(simp add: zeros.simps) + have s1: "\st i \* ps ia \* one ia \* ones (ia + 1) ja \* zero (ja + 1)\ ?body + \st i \* ps (ia + 1) \* zero ia \* ones (ia + 1) ja \* zero (ja + 1)\" + proof(cases "ja < ia + 1") + case True + from True False have "ja = ia" by auto + thus ?thesis + apply(simp add: ones.simps) + by (hsteps) + next + case False + from False have "ones (ia + 1) ja = (one (ia + 1) \* ones (ia + 1 + 1) ja)" + by(simp add: ones.simps) + thus ?thesis + by (simp, hsteps) + qed + have s2: "\st i \* ps (ia + 1) \* zero ia \* ones (ia + 1) ja \* zero (ja + 1)\ + ?body + \st j \* ps (ja + 1) \* zero ia \* zeros (ia + 1) ja \* zero (ja + 1)\" + by hsteps + from tm.sequencing[OF s1 s2] have + "\st i \* ps ia \* one ia \* ones (ia + 1) ja \* zero (ja + 1)\ ?body + \st j \* ps (ja + 1) \* zero ia \* zeros (ia + 1) ja \* zero (ja + 1)\" . + thus ?thesis + unfolding ones_eq zeros_eq by(simp add: sep_conj_ac) + qed + qed +qed + +lemma hoare_clear_until_zero_gen[step]: + assumes "u = v" "x = w + 1" + shows "\st i ** ps u ** ones v w ** zero x\ + i :[clear_until_zero]: j + \st j ** ps x ** zeros v w ** zero x\" + by (unfold assms, rule hoare_clear_until_zero) + +declare ones_simps[simp del] + +lemma hoare_move_left_reps[step]: + assumes "ks \ []" "u = v" + shows + "\st i ** ps u ** reps v w ks\ + i:[move_left]:j + \st j ** ps (u - 1) ** reps v w ks\" +proof - + from `ks \ []` obtain y ys where eq_ks: "ks = y#ys" + by (metis neq_Nil_conv) + show ?thesis + apply (unfold assms eq_ks) + apply (case_tac ys, simp) + my_block + have "(ones v (v + int y)) = (one v \* ones (v + 1) (v + int y))" + by (smt ones_step_simp) + my_block_end + apply (unfold this, hsteps) + by (simp add:this, hsteps) +qed + +lemma hoare_shift_left_cons: + assumes h: "ks \ []" + shows "\st i \* ps u \* tm (u - 1) x \* reps u v ks \* zero (v + 1) \* zero (v + 2) \ + i:[shift_left]:j + \st j \* ps (v + 2) \* reps (u - 1) (v - 1) ks \* zero v \* zero (v + 1) \* zero (v + 2) \" +proof(unfold shift_left_def, intro t_hoare_local t_hoare_label, clarify, + rule t_hoare_label_last, simp+, clarify, prune) + show " \st i \* ps u \* tm (u - 1) x \* reps u v ks \* zero (v + 1) \* zero (v + 2)\ + i :[ (if_zero j ; + move_left ; + write_one ; + right_until_zero ; + move_left ; write_zero ; + move_right ; move_right ; jmp i) ]: j + \st j \* ps (v + 2) \* reps (u - 1) (v - 1) ks \* zero v \* zero (v + 1) \* zero (v + 2)\" + using h + proof(induct ks arbitrary:i u v x) + case (Cons k ks) + thus ?case + proof(cases "ks = []") + let ?body = "i :[ (if_zero j ; move_left ; write_one ; right_until_zero ; + move_left ; write_zero ; move_right ; move_right ; jmp i) ]: j" + case True + have "\st i \* ps u \* tm (u - 1) x \* (one u \* ones (u + 1) (u + int k)) \* + zero (u + int k + 1) \* zero (u + int k + 2)\ + ?body + \st j \* ps (u + int k + 2) \* (one (u - 1) \* ones u (u - 1 + int k)) \* + zero (u + int k) \* zero (u + int k + 1) \* zero (u + int k + 2)\" + apply(rule tm.sequencing [where q = "st i \* ps (u + int k + 2) \* + (one (u - 1) \* ones u (u - 1 + int k)) \* + zero (u + int k) \* zero (u + int k + 1) \* zero (u + int k + 2)"]) + apply (hsteps) + apply (rule_tac p = "st j' \* ps (u - 1) \* ones (u - 1) (u + int k) \* + zero (u + int k + 1) \* zero (u + int k + 2)" + in tm.pre_stren) + apply (hsteps) + my_block + have "(ones (u - 1) (u + int k)) = (ones (u - 1) (u + int k - 1) \* one (u + int k))" + by (smt ones_rev) + my_block_end + apply (unfold this) + apply hsteps + apply (simp add:sep_conj_ac, sep_cancel+) + apply (smt ones.simps sep.mult_assoc sep_conj_commuteI) + apply (simp add:sep_conj_ac)+ + apply (sep_cancel+) + apply (smt ones.simps sep.mult_left_commute sep_conj_commuteI this) + by hstep + with True show ?thesis + by (simp add:ones_simps, simp only:sep_conj_cond, intro tm.pre_condI, simp) + next + case False + let ?body = "i :[ (if_zero j ; move_left ; write_one ;right_until_zero ; move_left ; + write_zero ; move_right ; move_right ; jmp i) ]: j" + have "\st i \* ps u \* tm (u - 1) x \* one u \* ones (u + 1) (u + int k) \* + zero (u + int k + 1) \* reps (u + int k + 2) v ks \* zero (v + 1) \* zero (v + 2)\ + ?body + \st j \* ps (v + 2) \* one (u - 1) \* ones u (u - 1 + int k) \* + zero (u + int k) \* reps (1 + (u + int k)) (v - 1) ks \* + zero v \* zero (v + 1) \* zero (v + 2)\" + apply (rule tm.sequencing[where q = "st i \* ps (u + int k + 2) \* + zero (u + int k + 1) \* reps (u + int k + 2) v ks \* zero (v + 1) \* + zero (v + 2) \* one (u - 1) \* ones u (u - 1 + int k) \* zero (u + int k)"]) + apply (hsteps) + apply (rule_tac p = "st j' \* ps (u - 1) \* + ones (u - 1) (u + int k) \* + zero (u + int k + 1) \* + reps (u + int k + 2) v ks \* zero (v + 1) \* zero (v + 2) + " in tm.pre_stren) + apply hsteps + my_block + have "(ones (u - 1) (u + int k)) = + (ones (u - 1) (u + int k - 1) \* one (u + int k))" + by (smt ones_rev) + my_block_end + apply (unfold this) + apply (hsteps) + apply (sep_cancel+) + apply (smt ones.simps sep.mult_assoc sep_conj_commuteI) + apply (sep_cancel+) + apply (smt ones.simps this) + my_block + have eq_u: "1 + (u + int k) = u + int k + 1" by simp + from Cons.hyps[OF `ks \ []`, of i "u + int k + 2" Bk v, folded zero_def] + have "\st i \* ps (u + int k + 2) \* zero (u + int k + 1) \* + reps (u + int k + 2) v ks \* zero (v + 1) \* zero (v + 2)\ + ?body + \st j \* ps (v + 2) \* reps (1 + (u + int k)) (v - 1) ks \* + zero v \* zero (v + 1) \* zero (v + 2)\" + by (simp add:eq_u) + my_block_end my_note hh[step] = this + by hsteps + thus ?thesis + by (unfold reps_simp3[OF False], auto simp:sep_conj_ac ones_simps) + qed + qed auto +qed + +lemma hoare_shift_left_cons_gen[step]: + assumes h: "ks \ []" + "v = u - 1" "w = u" "y = x + 1" "z = x + 2" + shows "\st i \* ps u \* tm v vv \* reps w x ks \* tm y Bk \* tm z Bk\ + i:[shift_left]:j + \st j \* ps z \* reps v (x - 1) ks \* zero x \* zero y \* zero z \" + by (unfold assms, fold zero_def, rule hoare_shift_left_cons[OF `ks \ []`]) + +lemma hoare_bone_1_out: + assumes h: + "\ i j . \st i \* ps u \* zero u \* p \ + i:[c1]:j + \st e \* q \ + " + shows "\st i \* ps u \* zero u \* p \ + i:[(bone c1 c2)]:j + \st e \* q \ + " +apply (unfold bone_def, intro t_hoare_local) +apply hsteps +apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension) +apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension) +by (rule h) + +lemma hoare_bone_1: + assumes h: + "\ i j . \st i \* ps u \* zero u \* p \ + i:[c1]:j + \st j \* ps v \* tm v x \* q \ + " + shows "\st i \* ps u \* zero u \* p \ + i:[(bone c1 c2)]:j + \st j \* ps v \* tm v x \* q \ + " +proof - + note h[step] + show ?thesis + apply (unfold bone_def, intro t_hoare_local) + apply (rule t_hoare_label_last, auto) + apply hsteps + apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension) + by hsteps +qed + +lemma hoare_bone_2: + assumes h: + "\ i j . \st i \* ps u \* one u \* p \ + i:[c2]:j + \st j \* q \ + " + shows "\st i \* ps u \* one u \* p \ + i:[(bone c1 c2)]:j + \st j \* q \ + " +apply (unfold bone_def, intro t_hoare_local) +apply (rule_tac q = "st la \* ps u \* one u \* p" in tm.sequencing) +apply hsteps +apply (subst tassemble_to.simps(2), intro tm.code_exI, intro tm.code_extension1) +apply (subst tassemble_to.simps(2), intro tm.code_exI, intro tm.code_extension1) +apply (subst tassemble_to.simps(2), intro tm.code_exI) +apply (subst tassemble_to.simps(4), intro tm.code_condI, simp) +apply (subst tassemble_to.simps(2), intro tm.code_exI) +apply (subst tassemble_to.simps(4), simp add:sep_conj_cond, rule tm.code_condI, simp ) +by (rule h) + +lemma hoare_bone_2_out: + assumes h: + "\ i j . \st i \* ps u \* one u \* p \ + i:[c2]:j + \st e \* q \ + " + shows "\st i \* ps u \* one u \* p \ + i:[(bone c1 c2)]:j + \st e \* q \ + " +apply (unfold bone_def, intro t_hoare_local) +apply (rule_tac q = "st la \* ps u \* one u \* p" in tm.sequencing) +apply hsteps +apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension1) +apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension1) +apply (subst tassemble_to.simps(2), intro tm.code_exI) +apply (subst tassemble_to.simps(4), rule tm.code_condI, simp) +apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension) +by (rule h) + +lemma hoare_bzero_1: + assumes h[step]: + "\ i j . \st i \* ps u \* one u \* p \ + i:[c1]:j + \st j \* ps v \* tm v x \* q \ + " + shows "\st i \* ps u \* one u \* p \ + i:[(bzero c1 c2)]:j + \st j \* ps v \* tm v x \* q \ + " +apply (unfold bzero_def, intro t_hoare_local) +apply hsteps +apply (rule_tac c = " ((c1 ; jmp l) ; TLabel la ; c2 ; TLabel l)" in t_hoare_label_last, auto) +apply (subst tassemble_to.simps(2), intro tm.code_exI, intro tm.code_extension) +by hsteps + +lemma hoare_bzero_1_out: + assumes h[step]: + "\ i j . \st i \* ps u \* one u \* p \ + i:[c1]:j + \st e \* q \ + " + shows "\st i \* ps u \* one u \* p \ + i:[(bzero c1 c2)]:j + \st e \* q \ + " +apply (unfold bzero_def, intro t_hoare_local) +apply hsteps +apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension) +apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension) +by (rule h) + +lemma hoare_bzero_2: + assumes h: + "\ i j. \st i \* ps u \* zero u \* p \ + i:[c2]:j + \st j \* q \ + " + shows "\st i \* ps u \* zero u \* p \ + i:[(bzero c1 c2)]:j + \st j \* q \ + " + apply (unfold bzero_def, intro t_hoare_local) + apply (rule_tac q = "st la \* ps u \* zero u \* p" in tm.sequencing) + apply hsteps + apply (subst tassemble_to.simps(2), intro tm.code_exI, intro tm.code_extension1) + apply (subst tassemble_to.simps(2), intro tm.code_exI, intro tm.code_extension1) + apply (subst tassemble_to.simps(2), intro tm.code_exI) + apply (subst tassemble_to.simps(4)) + apply (rule tm.code_condI, simp) + apply (subst tassemble_to.simps(2)) + apply (rule tm.code_exI) + apply (subst tassemble_to.simps(4), simp add:sep_conj_cond) + apply (rule tm.code_condI, simp) + by (rule h) + +lemma hoare_bzero_2_out: + assumes h: + "\ i j . \st i \* ps u \* zero u \* p \ + i:[c2]:j + \st e \* q \ + " + shows "\st i \* ps u \* zero u \* p\ + i:[(bzero c1 c2)]:j + \st e \* q \ + " +apply (unfold bzero_def, intro t_hoare_local) +apply (rule_tac q = "st la \* ps u \* zero u \* p" in tm.sequencing) +apply hsteps +apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension1) +apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension1) +apply (subst tassemble_to.simps(2), intro tm.code_exI) +apply (subst tassemble_to.simps(4), rule tm.code_condI, simp) +apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension) +by (rule h) + +lemma reps_len_split: + assumes "xs \ []" "ys \ []" + shows "reps_len (xs @ ys) = reps_len xs + reps_len ys + 1" + using assms +proof(induct xs arbitrary:ys) + case (Cons x1 xs1) + show ?case + proof(cases "xs1 = []") + case True + thus ?thesis + by (simp add:reps_len_cons[OF `ys \ []`] reps_len_sg) + next + case False + hence " xs1 @ ys \ []" by simp + thus ?thesis + apply (simp add:reps_len_cons[OF `xs1@ys \ []`] reps_len_cons[OF `xs1 \ []`]) + by (simp add: Cons.hyps[OF `xs1 \ []` `ys \ []`]) + qed +qed auto + +lemma hoare_skip_or_set_set: + "\ st i \* ps u \* zero u \* zero (u + 1) \* tm (u + 2) x\ + i:[skip_or_set]:j + \ st j \* ps (u + 2) \* one u \* zero (u + 1) \* tm (u + 2) x\" + apply(unfold skip_or_set_def) + apply(rule_tac q = "st j \* ps (u + 2) \* tm (u + 2) x \* one u \* zero (u + 1)" + in tm.post_weaken) + apply(rule hoare_bone_1) + apply hsteps + by (auto simp:sep_conj_ac, sep_cancel+, smt) + +lemma hoare_skip_or_set_set_gen[step]: + assumes "u = v" "w = v + 1" "x = v + 2" + shows "\st i \* ps u \* zero v \* zero w \* tm x xv\ + i:[skip_or_set]:j + \st j \* ps x \* one v \* zero w \* tm x xv\" + by (unfold assms, rule hoare_skip_or_set_set) + +lemma hoare_skip_or_set_skip: + "\ st i \* ps u \* reps u v [k] \* zero (v + 1)\ + i:[skip_or_set]:j + \ st j \* ps (v + 2) \* reps u v [k] \* zero (v + 1)\" +proof - + show ?thesis + apply(unfold skip_or_set_def, unfold reps.simps, simp add:sep_conj_cond) + apply(rule tm.pre_condI, simp) + apply(rule_tac p = "st i \* ps u \* one u \* ones (u + 1) (u + int k) \* + zero (u + int k + 1)" + in tm.pre_stren) + apply (rule_tac q = "st j \* ps (u + int k + 2) \* + one u \* ones (u + 1) (u + int k) \* zero (u + int k + 1) + " in tm.post_weaken) + apply (rule hoare_bone_2) + apply (rule_tac p = " st i \* ps u \* ones u (u + int k) \* zero (u + int k + 1) + " in tm.pre_stren) + apply hsteps + apply (simp add:sep_conj_ac, sep_cancel+, auto simp:sep_conj_ac ones_simps) + by (sep_cancel+, smt) + qed + +lemma hoare_skip_or_set_skip_gen[step]: + assumes "u = v" "x = w + 1" + shows "\ st i \* ps u \* reps v w [k] \* zero x\ + i:[skip_or_set]:j + \ st j \* ps (w + 2) \* reps v w [k] \* zero x\" + by (unfold assms, rule hoare_skip_or_set_skip) + +lemma hoare_if_reps_z_true: + assumes h: "k = 0" + shows + "\st i \* ps u \* reps u v [k] \* zero (v + 1)\ + i:[if_reps_z e]:j + \st e \* ps u \* reps u v [k] \* zero (v + 1)\" + apply (unfold reps.simps, simp add:sep_conj_cond) + apply (rule tm.pre_condI, simp add:h) + apply (unfold if_reps_z_def) + apply (simp add:ones_simps) + apply (hsteps) + apply (rule_tac p = "st j' \* ps (u + 1) \* zero (u + 1) \* one u" in tm.pre_stren) + apply (rule hoare_bone_1_out) + by (hsteps) + +lemma hoare_if_reps_z_true_gen[step]: + assumes "k = 0" "u = v" "x = w + 1" + shows "\st i \* ps u \* reps v w [k] \* zero x\ + i:[if_reps_z e]:j + \st e \* ps u \* reps v w [k] \* zero x\" + by (unfold assms, rule hoare_if_reps_z_true, simp) + +lemma hoare_if_reps_z_false: + assumes h: "k \ 0" + shows + "\st i \* ps u \* reps u v [k]\ + i:[if_reps_z e]:j + \st j \* ps u \* reps u v [k]\" +proof - + from h obtain k' where "k = Suc k'" by (metis not0_implies_Suc) + show ?thesis + apply (unfold `k = Suc k'`) + apply (simp add:sep_conj_cond, rule tm.pre_condI, simp) + apply (unfold if_reps_z_def) + apply (simp add:ones_simps) + apply hsteps + apply (rule_tac p = "st j' \* ps (u + 1) \* one (u + 1) \* one u \* + ones (2 + u) (u + (1 + int k'))" in tm.pre_stren) + apply (rule_tac hoare_bone_2) + by (hsteps) +qed + +lemma hoare_if_reps_z_false_gen[step]: + assumes h: "k \ 0" "u = v" + shows + "\st i \* ps u \* reps v w [k]\ + i:[if_reps_z e]:j + \st j \* ps u \* reps v w [k]\" + by (unfold assms, rule hoare_if_reps_z_false[OF `k \ 0`]) + +lemma EXS_postI: + assumes "\P\ + c + \Q x\" + shows "\P\ + c + \EXS x. Q x\" +by (metis EXS_intro assms tm.hoare_adjust) + +lemma hoare_if_reps_nz_true: + assumes h: "k \ 0" + shows + "\st i \* ps u \* reps u v [k]\ + i:[if_reps_nz e]:j + \st e \* ps u \* reps u v [k]\" +proof - + from h obtain k' where "k = Suc k'" by (metis not0_implies_Suc) + show ?thesis + apply (unfold `k = Suc k'`) + apply (unfold reps.simps, simp add:sep_conj_cond, rule tm.pre_condI, simp) + apply (unfold if_reps_nz_def) + apply (simp add:ones_simps) + apply hsteps + apply (rule_tac p = "st j' \* ps (u + 1) \* one (u + 1) \* one u \* + ones (2 + u) (u + (1 + int k'))" in tm.pre_stren) + apply (rule hoare_bzero_1_out) + by hsteps +qed + + +lemma hoare_if_reps_nz_true_gen[step]: + assumes h: "k \ 0" "u = v" + shows + "\st i \* ps u \* reps v w [k]\ + i:[if_reps_nz e]:j + \st e \* ps u \* reps v w [k]\" + by (unfold assms, rule hoare_if_reps_nz_true[OF `k\ 0`]) + +lemma hoare_if_reps_nz_false: + assumes h: "k = 0" + shows + "\st i \* ps u \* reps u v [k] \* zero (v + 1)\ + i:[if_reps_nz e]:j + \st j \* ps u \* reps u v [k] \* zero (v + 1)\" + apply (simp add:h sep_conj_cond) + apply (rule tm.pre_condI, simp) + apply (unfold if_reps_nz_def) + apply (simp add:ones_simps) + apply (hsteps) + apply (rule_tac p = "st j' \* ps (u + 1) \* zero (u + 1) \* one u" in tm.pre_stren) + apply (rule hoare_bzero_2) + by (hsteps) + +lemma hoare_if_reps_nz_false_gen[step]: + assumes h: "k = 0" "u = v" "x = w + 1" + shows + "\st i \* ps u \* reps v w [k] \* zero x\ + i:[if_reps_nz e]:j + \st j \* ps u \* reps v w [k] \* zero x\" + by (unfold assms, rule hoare_if_reps_nz_false, simp) + +lemma hoare_skip_or_sets_set: + shows "\st i \* ps u \* zeros u (u + int (reps_len (replicate (Suc n) 0))) \* + tm (u + int (reps_len (replicate (Suc n) 0)) + 1) x\ + i:[skip_or_sets (Suc n)]:j + \st j \* ps (u + int (reps_len (replicate (Suc n) 0)) + 1) \* + reps' u (u + int (reps_len (replicate (Suc n) 0))) (replicate (Suc n) 0) \* + tm (u + int (reps_len (replicate (Suc n) 0)) + 1) x \" +proof(induct n arbitrary:i j u x) + case 0 + from 0 show ?case + apply (simp add:reps'_def reps_len_def reps_ctnt_len_def reps_sep_len_def reps.simps) + apply (unfold skip_or_sets_def, simp add:tpg_fold_sg) + apply hsteps + by (auto simp:sep_conj_ac, smt cond_true_eq2 ones.simps sep_conj_left_commute) +next + case (Suc n) + { fix n + have "listsum (replicate n (Suc 0)) = n" + by (induct n, auto) + } note eq_sum = this + have eq_len: "\n. n \ 0 \ reps_len (replicate (Suc n) 0) = reps_len (replicate n 0) + 2" + by (simp add:reps_len_def reps_sep_len_def reps_ctnt_len_def) + have eq_zero: "\ u v. (zeros u (u + int (v + 2))) = + (zeros u (u + (int v)) \* zero (u + (int v) + 1) \* zero (u + (int v) + 2))" + by (smt sep.mult_assoc zeros_rev) + hence eq_z: + "zeros u (u + int (reps_len (replicate (Suc (Suc n)) 0))) = + (zeros u (u + int (reps_len (replicate (Suc n) 0))) \* + zero ((u + int (reps_len (replicate (Suc n) 0))) + 1) \* + zero ((u + int (reps_len (replicate (Suc n) 0))) + 2)) + " by (simp only:eq_len) + have hh: "\x. (replicate (Suc (Suc n)) x) = (replicate (Suc n) x) @ [x]" + by (metis replicate_Suc replicate_append_same) + have hhh: "replicate (Suc n) skip_or_set \ []" "[skip_or_set] \ []" by auto + have eq_code: + "(i :[ skip_or_sets (Suc (Suc n)) ]: j) = + (i :[ (skip_or_sets (Suc n); skip_or_set) ]: j)" + proof(unfold skip_or_sets_def) + show "i :[ tpg_fold (replicate (Suc (Suc n)) skip_or_set) ]: j = + i :[ (tpg_fold (replicate (Suc n) skip_or_set) ; skip_or_set) ]: j" + apply (insert tpg_fold_app[OF hhh, of i j], unfold hh) + by (simp only:tpg_fold_sg) + qed + have "Suc n \ 0" by simp + show ?case + apply (unfold eq_z eq_code) + apply (hstep Suc(1)) + apply (unfold eq_len[OF `Suc n \ 0`]) + apply hstep + apply (auto simp:sep_conj_ac)[1] + apply (sep_cancel+, prune) + apply (fwd abs_ones) + apply ((fwd abs_reps')+, simp add:int_add_ac) + by (metis replicate_append_same) + qed + +lemma hoare_skip_or_sets_set_gen[step]: + assumes h: "p2 = p1" + "p3 = p1 + int (reps_len (replicate (Suc n) 0))" + "p4 = p3 + 1" + shows "\st i \* ps p1 \* zeros p2 p3 \* tm p4 x\ + i:[skip_or_sets (Suc n)]:j + \st j \* ps p4 \* reps' p2 p3 (replicate (Suc n) 0) \* tm p4 x\" + apply (unfold h) + by (rule hoare_skip_or_sets_set) + +declare reps.simps[simp del] + +lemma hoare_skip_or_sets_skip: + assumes h: "n < length ks" + shows "\st i \* ps u \* reps' u v (take n ks) \* reps' (v + 1) w [ks!n] \ + i:[skip_or_sets (Suc n)]:j + \st j \* ps (w+1) \* reps' u v (take n ks) \* reps' (v + 1) w [ks!n]\" + using h +proof(induct n arbitrary: i j u v w ks) + case 0 + show ?case + apply (subst (1 5) reps'_def, simp add:sep_conj_cond, intro tm.pre_condI, simp) + apply (unfold skip_or_sets_def, simp add:tpg_fold_sg) + apply (unfold reps'_def, simp del:reps.simps) + apply hsteps + by (sep_cancel+, smt+) +next + case (Suc n) + from `Suc n < length ks` have "n < length ks" by auto + note h = Suc(1) [OF this] + show ?case + my_block + from `Suc n < length ks` + have eq_take: "take (Suc n) ks = take n ks @ [ks!n]" + by (metis not_less_eq not_less_iff_gr_or_eq take_Suc_conv_app_nth) + my_block_end + apply (unfold this) + apply (subst reps'_append, simp add:sep_conj_exists, rule tm.precond_exI) + my_block + have "(i :[ skip_or_sets (Suc (Suc n)) ]: j) = + (i :[ (skip_or_sets (Suc n); skip_or_set )]: j)" + proof - + have eq_rep: + "(replicate (Suc (Suc n)) skip_or_set) = ((replicate (Suc n) skip_or_set) @ [skip_or_set])" + by (metis replicate_Suc replicate_append_same) + have "replicate (Suc n) skip_or_set \ []" "[skip_or_set] \ []" by auto + from tpg_fold_app[OF this] + show ?thesis + by (unfold skip_or_sets_def eq_rep, simp del:replicate.simps add:tpg_fold_sg) + qed + my_block_end + apply (unfold this) + my_block + fix i j m + have "\st i \* ps u \* (reps' u (m - 1) (take n ks) \* reps' m v [ks ! n])\ + i :[ (skip_or_sets (Suc n)) ]: j + \st j \* ps (v + 1) \* (reps' u (m - 1) (take n ks) \* reps' m v [ks ! n])\" + apply (rule h[THEN tm.hoare_adjust]) + by (sep_cancel+, auto) + my_block_end my_note h_c1 = this + my_block + fix j' j m + have "\st j' \* ps (v + 1) \* reps' (v + 1) w [ks ! Suc n]\ + j' :[ skip_or_set ]: j + \st j \* ps (w + 1) \* reps' (v + 1) w [ks ! Suc n]\" + apply (unfold reps'_def, simp) + apply (rule hoare_skip_or_set_skip[THEN tm.hoare_adjust]) + by (sep_cancel+, smt)+ + my_block_end + apply (hstep h_c1 this)+ + by ((fwd abs_reps'), simp, sep_cancel+) +qed + +lemma hoare_skip_or_sets_skip_gen[step]: + assumes h: "n < length ks" "u = v" "x = w + 1" + shows "\st i \* ps u \* reps' v w (take n ks) \* reps' x y [ks!n] \ + i:[skip_or_sets (Suc n)]:j + \st j \* ps (y+1) \* reps' v w (take n ks) \* reps' x y [ks!n]\" + by (unfold assms, rule hoare_skip_or_sets_skip[OF `n < length ks`]) + +lemma fam_conj_interv_simp: + "(fam_conj {(ia::int)<..} p) = ((p (ia + 1)) \* fam_conj {ia + 1 <..} p)" +by (smt Collect_cong fam_conj_insert_simp greaterThan_def + greaterThan_eq_iff greaterThan_iff insertI1 + insert_compr lessThan_iff mem_Collect_eq) + +lemma zeros_fam_conj: + assumes "u \ v" + shows "(zeros u v \* fam_conj {v<..} zero) = fam_conj {u - 1<..} zero" +proof - + have "{u - 1<..v} ## {v <..}" by (auto simp:set_ins_def) + from fam_conj_disj_simp[OF this, symmetric] + have "(fam_conj {u - 1<..v} zero \* fam_conj {v<..} zero) = fam_conj ({u - 1<..v} + {v<..}) zero" . + moreover + from `u \ v` have eq_set: "{u - 1 <..} = {u - 1 <..v} + {v <..}" by (auto simp:set_ins_def) + moreover have "fam_conj {u - 1<..v} zero = zeros u v" + proof - + have "({u - 1<..v}) = ({u .. v})" by auto + moreover { + fix u v + assume "u \ (v::int)" + hence "fam_conj {u .. v} zero = zeros u v" + proof(induct rule:ones_induct) + case (Base i j) + thus ?case by auto + next + case (Step i j) + thus ?case + proof(cases "i = j") + case True + show ?thesis + by (unfold True, simp add:fam_conj_simps) + next + case False + with `i \ j` have hh: "i + 1 \ j" by auto + hence eq_set: "{i..j} = (insert i {i + 1 .. j})" + by (smt simp_from_to) + have "i \ {i + 1 .. j}" by simp + from fam_conj_insert_simp[OF this, folded eq_set] + have "fam_conj {i..j} zero = (zero i \* fam_conj {i + 1..j} zero)" . + with Step(2)[OF hh] Step + show ?thesis by simp + qed + qed + } + moreover note this[OF `u \ v`] + ultimately show ?thesis by simp + qed + ultimately show ?thesis by smt +qed + +declare replicate.simps [simp del] + +lemma hoare_skip_or_sets_comb: + assumes "length ks \ n" + shows "\st i \* ps u \* reps u v ks \* fam_conj {v<..} zero\ + i:[skip_or_sets (Suc n)]:j + \st j \* ps ((v + int (reps_len (list_ext n ks)) - int (reps_len ks))+ 2) \* + reps' u (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1) (list_ext n ks) \* + fam_conj {(v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1)<..} zero \" +proof(cases "ks = []") + case True + show ?thesis + apply (subst True, simp only:reps.simps sep_conj_cond) + apply (rule tm.pre_condI, simp) + apply (rule_tac p = "st i \* ps (v + 1) \* + zeros (v + 1) (v + 1 + int (reps_len (replicate (Suc n) 0))) \* + tm (v + 2 + int (reps_len (replicate (Suc n) 0))) Bk \* + fam_conj {(v + 2 + int (reps_len (replicate (Suc n) 0)))<..} zero + " in tm.pre_stren) + apply hsteps + apply (auto simp:sep_conj_ac)[1] + apply (auto simp:sep_conj_ac)[2] + my_block + from True have "(list_ext n ks) = (replicate (Suc n) 0)" + by (metis append_Nil diff_zero list.size(3) list_ext_def) + my_block_end my_note le_red = this + my_block + from True have "(reps_len ks) = 0" + by (metis reps_len_nil) + my_block_end + apply (unfold this le_red, simp) + my_block + have "v + 2 + int (reps_len (replicate (Suc n) 0)) = + v + int (reps_len (replicate (Suc n) 0)) + 2" by smt + my_block_end my_note eq_len = this + apply (unfold this) + apply (sep_cancel+) + apply (fold zero_def) + apply (subst fam_conj_interv_simp, simp) + apply (simp only:int_add_ac) + apply (simp only:sep_conj_ac, sep_cancel+) + my_block + have "v + 1 \ (2 + (v + int (reps_len (replicate (Suc n) 0))))" by simp + from zeros_fam_conj[OF this] + have "(fam_conj {v<..} zero) = (zeros (v + 1) (2 + (v + int (reps_len (replicate (Suc n) 0)))) \* + fam_conj {2 + (v + int (reps_len (replicate (Suc n) 0)))<..} zero)" + by simp + my_block_end + apply (subst (asm) this, simp only:int_add_ac, sep_cancel+) + by (smt cond_true_eq2 sep.mult_assoc sep.mult_commute + sep.mult_left_commute sep_conj_assoc sep_conj_commute + sep_conj_left_commute zeros.simps zeros_rev) +next + case False + show ?thesis + my_block + have "(i:[skip_or_sets (Suc n)]:j) = + (i:[(skip_or_sets (length ks); skip_or_sets (Suc n - length ks))]:j)" + apply (unfold skip_or_sets_def) + my_block + have "(replicate (Suc n) skip_or_set) = + (replicate (length ks) skip_or_set @ (replicate (Suc n - length ks) skip_or_set))" + by (smt assms replicate_add) + my_block_end + apply (unfold this, rule tpg_fold_app, simp add:False) + by (insert `length ks \ n`, simp) + my_block_end + apply (unfold this) + my_block + from False have "length ks = (Suc (length ks - 1))" by simp + my_block_end + apply (subst (1) this) + my_block + from False + have "(reps u v ks \* fam_conj {v<..} zero) = + (reps' u (v + 1) ks \* fam_conj {v+1<..} zero)" + apply (unfold reps'_def, simp) + by (subst fam_conj_interv_simp, simp add:sep_conj_ac) + my_block_end + apply (unfold this) + my_block + fix i j + have "\st i \* ps u \* reps' u (v + 1) ks \ + i :[ skip_or_sets (Suc (length ks - 1))]: j + \st j \* ps (v + 2) \* reps' u (v + 1) ks \" + my_block + have "ks = take (length ks - 1) ks @ [ks!(length ks - 1)]" + by (smt False drop_0 drop_eq_Nil id_take_nth_drop) + my_block_end my_note eq_ks = this + apply (subst (1) this) + apply (unfold reps'_append, simp add:sep_conj_exists, rule tm.precond_exI) + my_block + from False have "(length ks - Suc 0) < length ks" + by (smt `length ks = Suc (length ks - 1)`) + my_block_end + apply hsteps + apply (subst eq_ks, unfold reps'_append, simp only:sep_conj_exists) + by (rule_tac x = m in EXS_intro, simp add:sep_conj_ac, sep_cancel+, smt) + my_block_end + apply (hstep this) + my_block + fix u n + have "(fam_conj {u <..} zero) = + (zeros (u + 1) (u + int n + 1) \* tm (u + int n + 2) Bk \* fam_conj {(u + int n + 2)<..} zero)" + my_block + have "u + 1 \ (u + int n + 2)" by auto + from zeros_fam_conj[OF this, symmetric] + have "fam_conj {u<..} zero = (zeros (u + 1) (u + int n + 2) \* fam_conj {u + int n + 2<..} zero)" + by simp + my_block_end + apply (subst this) + my_block + have "(zeros (u + 1) (u + int n + 2)) = + ((zeros (u + 1) (u + int n + 1) \* tm (u + int n + 2) Bk))" + by (smt zero_def zeros_rev) + my_block_end + by (unfold this, auto simp:sep_conj_ac) + my_block_end + apply (subst (1) this[of _ "(reps_len (replicate (Suc (n - length ks)) 0))"]) + my_block + from `length ks \ n` + have "Suc n - length ks = Suc (n - length ks)" by auto + my_block_end my_note eq_suc = this + apply (subst this) + apply hsteps + apply (simp add: sep_conj_ac this, sep_cancel+) + apply (fwd abs_reps')+ + my_block + have "(int (reps_len (replicate (Suc (n - length ks)) 0))) = + (int (reps_len (list_ext n ks)) - int (reps_len ks) - 1)" + apply (unfold list_ext_def eq_suc) + my_block + have "replicate (Suc (n - length ks)) 0 \ []" by simp + my_block_end + by (unfold reps_len_split[OF False this], simp) + my_block_end + apply (unfold this) + my_block + from `length ks \ n` + have "(ks @ replicate (Suc (n - length ks)) 0) = (list_ext n ks)" + by (unfold list_ext_def, simp) + my_block_end + apply (unfold this, simp) + apply (subst fam_conj_interv_simp, unfold zero_def, simp, simp add:int_add_ac sep_conj_ac) + by (sep_cancel+, smt) +qed + +lemma hoare_skip_or_sets_comb_gen: + assumes "length ks \ n" "u = v" "w = x" + shows "\st i \* ps u \* reps v w ks \* fam_conj {x<..} zero\ + i:[skip_or_sets (Suc n)]:j + \st j \* ps ((x + int (reps_len (list_ext n ks)) - int (reps_len ks))+ 2) \* + reps' u (x + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1) (list_ext n ks) \* + fam_conj {(x + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1)<..} zero \" + by (unfold assms, rule hoare_skip_or_sets_comb[OF `length ks \ n`]) + +lemma list_ext_tail_expand: + assumes h: "length ks \ a" + shows "list_ext a ks = take a (list_ext a ks) @ [(list_ext a ks)!a]" +proof - + let ?l = "list_ext a ks" + from h have eq_len: "length ?l = Suc a" + by (smt list_ext_len_eq) + hence "?l \ []" by auto + hence "?l = take (length ?l - 1) ?l @ [?l!(length ?l - 1)]" + by (metis `length (list_ext a ks) = Suc a` diff_Suc_1 le_refl + lessI take_Suc_conv_app_nth take_all) + from this[unfolded eq_len] + show ?thesis by simp +qed + +lemma reps'_nn_expand: + assumes "xs \ []" + shows "(reps' u v xs) = (reps u (v - 1) xs \* zero v)" + by (metis assms reps'_def) + +lemma sep_conj_st1: "(p \* st t \* q) = (st t \* p \* q)" + by (simp only:sep_conj_ac) + +lemma sep_conj_st2: "(p \* st t) = (st t \* p)" + by (simp only:sep_conj_ac) + +lemma sep_conj_st3: "((st t \* p) \* r) = (st t \* p \* r)" + by (simp only:sep_conj_ac) + +lemma sep_conj_st4: "(EXS x. (st t \* r x)) = ((st t) \* (EXS x. r x))" + apply (unfold pred_ex_def, default+) + apply (safe) + apply (sep_cancel, auto) + by (auto elim!: sep_conjE intro!:sep_conjI) + +lemmas sep_conj_st = sep_conj_st1 sep_conj_st2 sep_conj_st3 sep_conj_st4 + +lemma sep_conj_cond3 : "( \* ) = <(cond1 \ cond2)>" + by (smt cond_eqI cond_true_eq sep_conj_commute sep_conj_empty) + +lemma sep_conj_cond4 : "( \* \* r) = (<(cond1 \ cond2)> \* r)" + by (metis Hoare_gen.sep_conj_cond3 Hoare_tm.sep_conj_cond3) + +lemmas sep_conj_cond = sep_conj_cond3 sep_conj_cond4 sep_conj_cond + +lemma hoare_left_until_zero_reps: + "\st i ** ps v ** zero u ** reps (u + 1) v [k]\ + i:[left_until_zero]:j + \st j ** ps u ** zero u ** reps (u + 1) v [k]\" + apply (unfold reps.simps, simp only:sep_conj_cond) + apply (rule tm.pre_condI, simp) + by hstep + +lemma hoare_left_until_zero_reps_gen[step]: + assumes "u = x" "w = v + 1" + shows "\st i ** ps u ** zero v ** reps w x [k]\ + i:[left_until_zero]:j + \st j ** ps v ** zero v ** reps w x [k]\" + by (unfold assms, rule hoare_left_until_zero_reps) + +lemma reps_lenE: + assumes "reps u v ks s" + shows "( <(v = u + int (reps_len ks) - 1)> \* reps u v ks ) s" +proof(rule condI) + from reps_len_correct[OF assms] show "v = u + int (reps_len ks) - 1" . +next + from assms show "reps u v ks s" . +qed + +lemma condI1: + assumes "p s" "b" + shows "( \* p) s" +proof(rule condI[OF assms(2)]) + from assms(1) show "p s" . +qed + +lemma hoare_locate_set: + assumes "length ks \ n" + shows "\st i \* zero (u - 1) \* ps u \* reps u v ks \* fam_conj {v<..} zero\ + i:[locate n]:j + \st j \* zero (u - 1) \* + (EXS m w. ps m \* reps' u (m - 1) (take n (list_ext n ks)) \* + reps m w [(list_ext n ks)!n] \* fam_conj {w<..} zero)\" +proof(cases "take n (list_ext n ks) = []") + case False + show ?thesis + apply (unfold locate_def) + apply (hstep hoare_skip_or_sets_comb_gen) + apply (subst (3) list_ext_tail_expand[OF `length ks \ n`]) + apply (subst (1) reps'_append, simp add:sep_conj_exists) + apply (rule tm.precond_exI) + apply (subst (1) reps'_nn_expand[OF False]) + apply (rule_tac p = "st j' \* <(m = u + int (reps_len (take n (list_ext n ks))) + 1)> \* + ps (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 2) \* + ((reps u (m - 1 - 1) (take n (list_ext n ks)) \* zero (m - 1)) \* + reps' m (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1) + [list_ext n ks ! n]) \* + fam_conj + {v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1<..} + zero \* + zero (u - 1)" + in tm.pre_stren) + my_block + have "[list_ext n ks ! n] \ []" by simp + from reps'_nn_expand[OF this] + have "(reps' m (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1) [list_ext n ks ! n]) = + (reps m (v + (int (reps_len (list_ext n ks)) - int (reps_len ks))) [list_ext n ks ! n] \* + zero (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1))" + by simp + my_block_end + apply (subst this) + my_block + have "(fam_conj {v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1<..} zero) = + (zero (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 2) \* + fam_conj {v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 2<..} zero)" + by (subst fam_conj_interv_simp, smt) + my_block_end + apply (unfold this) + apply (simp only:sep_conj_st) + apply hsteps + apply (auto simp:sep_conj_ac)[1] + apply (sep_cancel+) + apply (rule_tac x = m in EXS_intro) + apply (rule_tac x = "m + int (list_ext n ks ! n)" in EXS_intro) + apply (simp add:sep_conj_ac del:ones_simps, sep_cancel+) + apply (subst (asm) sep_conj_cond)+ + apply (erule_tac condE, clarsimp, simp add:sep_conj_ac int_add_ac) + apply (fwd abs_reps') + apply (fwd abs_reps') + apply (simp add:sep_conj_ac int_add_ac) + apply (sep_cancel+) + apply (subst (asm) reps'_def, subst fam_conj_interv_simp, subst fam_conj_interv_simp, + simp add:sep_conj_ac int_add_ac) + my_block + fix s + assume h: "(reps (1 + (u + int (reps_len (take n (list_ext n ks))))) + (v + (- int (reps_len ks) + int (reps_len (list_ext n ks)))) [list_ext n ks ! n]) s" + (is "?P s") + from reps_len_correct[OF this] list_ext_tail_expand[OF `length ks \ n`] + have hh: "v + (- int (reps_len ks) + + int (reps_len (take n (list_ext n ks) @ [list_ext n ks ! n]))) = + 1 + (u + int (reps_len (take n (list_ext n ks)))) + + int (reps_len [list_ext n ks ! n]) - 1" + by metis + have "[list_ext n ks ! n] \ []" by simp + from hh[unfolded reps_len_split[OF False this]] + have "v = u + (int (reps_len ks)) - 1" + by simp + from condI1[where p = ?P, OF h this] + have "(<(v = u + int (reps_len ks) - 1)> \* + reps (1 + (u + int (reps_len (take n (list_ext n ks))))) + (v + (- int (reps_len ks) + int (reps_len (list_ext n ks)))) [list_ext n ks ! n]) s" . + my_block_end + apply (fwd this, (subst (asm) sep_conj_cond)+, erule condE, simp add:sep_conj_ac int_add_ac + reps_len_sg) + apply (fwd reps_lenE, (subst (asm) sep_conj_cond)+, erule condE, simp add:sep_conj_ac int_add_ac + reps_len_sg) + by (fwd reps_lenE, (subst (asm) sep_conj_cond)+, erule condE, simp add:sep_conj_ac) +next + case True + show ?thesis + apply (unfold locate_def) + apply (hstep hoare_skip_or_sets_comb) + apply (subst (3) list_ext_tail_expand[OF `length ks \ n`]) + apply (subst (1) reps'_append, simp add:sep_conj_exists) + apply (rule tm.precond_exI) + my_block + from True `length ks \ n` + have "ks = []" "n = 0" + apply (metis le0 le_antisym length_0_conv less_nat_zero_code list_ext_len take_eq_Nil) + by (smt True length_take list.size(3) list_ext_len) + my_block_end + apply (unfold True this) + apply (simp add:reps'_def list_ext_def reps.simps reps_len_def reps_sep_len_def + reps_ctnt_len_def + del:ones_simps) + apply (subst sep_conj_cond)+ + apply (rule tm.pre_condI, simp del:ones_simps) + apply (subst fam_conj_interv_simp, simp add:sep_conj_st del:ones_simps) + apply (hsteps) + apply (auto simp:sep_conj_ac)[1] + apply (sep_cancel+) + apply (rule_tac x = "(v + int (listsum (replicate (Suc 0) (Suc 0))))" in EXS_intro)+ + apply (simp only:sep_conj_ac, sep_cancel+) + apply (auto) + apply (subst fam_conj_interv_simp) + apply (subst fam_conj_interv_simp) + by smt +qed + +lemma hoare_locate_set_gen[step]: + assumes "length ks \ n" + "u = v - 1" "v = w" "x = y" + shows "\st i \* zero u \* ps v \* reps w x ks \* fam_conj {y<..} zero\ + i:[locate n]:j + \st j \* zero u \* + (EXS m w. ps m \* reps' v (m - 1) (take n (list_ext n ks)) \* + reps m w [(list_ext n ks)!n] \* fam_conj {w<..} zero)\" + by (unfold assms, rule hoare_locate_set[OF `length ks \ n`]) + +lemma hoare_locate_skip: + assumes h: "n < length ks" + shows + "\st i \* ps u \* zero (u - 1) \* reps' u (v - 1) (take n ks) \* reps' v w [ks!n] \* tm (w + 1) x\ + i:[locate n]:j + \st j \* ps v \* zero (u - 1) \* reps' u (v - 1) (take n ks) \* reps' v w [ks!n] \* tm (w + 1) x\" +proof - + show ?thesis + apply (unfold locate_def) + apply hsteps + apply (subst (2 4) reps'_def, simp add:reps.simps sep_conj_cond del:ones_simps) + apply (intro tm.pre_condI, simp del:ones_simps) + apply hsteps + apply (case_tac "(take n ks) = []", simp add:reps'_def sep_conj_cond del:ones_simps) + apply (rule tm.pre_condI, simp del:ones_simps) + apply hsteps + apply (simp del:ones_simps add:reps'_def) + by hsteps +qed + + +lemma hoare_locate_skip_gen[step]: + assumes "n < length ks" + "v = u - 1" "w = u" "x = y - 1" "z' = z + 1" + shows + "\st i \* ps u \* tm v Bk \* reps' w x (take n ks) \* reps' y z [ks!n] \* tm z' vx\ + i:[locate n]:j + \st j \* ps y \* tm v Bk \* reps' w x (take n ks) \* reps' y z [ks!n] \* tm z' vx\" + by (unfold assms, fold zero_def, rule hoare_locate_skip[OF `n < length ks`]) + +lemma ones_int_expand: "(ones m (m + int k)) = (one m \* ones (m + 1) (m + int k))" + by (simp add:ones_simps) + +lemma reps_splitedI: + assumes h1: "(reps u v ks1 \* zero (v + 1) \* reps (v + 2) w ks2) s" + and h2: "ks1 \ []" + and h3: "ks2 \ []" + shows "(reps u w (ks1 @ ks2)) s" +proof - + from h2 h3 + have "splited (ks1 @ ks2) ks1 ks2" by (auto simp:splited_def) + from h1 obtain s1 where + "(reps u v ks1) s1" by (auto elim:sep_conjE) + from h1 obtain s2 where + "(reps (v + 2) w ks2) s2" by (auto elim:sep_conjE) + from reps_len_correct[OF `(reps u v ks1) s1`] + have eq_v: "v = u + int (reps_len ks1) - 1" . + from reps_len_correct[OF `(reps (v + 2) w ks2) s2`] + have eq_w: "w = v + 2 + int (reps_len ks2) - 1" . + from h1 + have "(reps u (u + int (reps_len ks1) - 1) ks1 \* + zero (u + int (reps_len ks1)) \* reps (u + int (reps_len ks1) + 1) w ks2) s" + apply (unfold eq_v eq_w[unfolded eq_v]) + by (sep_cancel+, smt) + thus ?thesis + by(unfold reps_splited[OF `splited (ks1 @ ks2) ks1 ks2`], simp) +qed + +lemma reps_sucI: + assumes h: "(reps u v (xs@[x]) \* one (1 + v)) s" + shows "(reps u (1 + v) (xs@[Suc x])) s" +proof(cases "xs = []") + case True + from h obtain s' where "(reps u v (xs@[x])) s'" by (auto elim:sep_conjE) + from reps_len_correct[OF this] have " v = u + int (reps_len (xs @ [x])) - 1" . + with True have eq_v: "v = u + int x" by (simp add:reps_len_sg) + have eq_one1: "(one (1 + (u + int x)) \* ones (u + 1) (u + int x)) = ones (u + 1) (1 + (u + int x))" + by (smt ones_rev sep.mult_commute) + from h show ?thesis + apply (unfold True, simp add:eq_v reps.simps sep_conj_cond sep_conj_ac ones_rev) + by (sep_cancel+, simp add: eq_one1, smt) +next + case False + from h obtain s1 s2 where hh: "(reps u v (xs@[x])) s1" "s = s1 + s2" "s1 ## s2" "one (1 + v) s2" + by (auto elim:sep_conjE) + from hh(1)[unfolded reps_rev[OF False]] + obtain s11 s12 s13 where hhh: + "(reps u (v - int (x + 1) - 1) xs) s11" + "(zero (v - int (x + 1))) s12" "(ones (v - int x) v) s13" + "s11 ## (s12 + s13)" "s12 ## s13" "s1 = s11 + s12 + s13" + apply (atomize_elim) + apply(elim sep_conjE)+ + apply (rule_tac x = xa in exI) + apply (rule_tac x = xaa in exI) + apply (rule_tac x = ya in exI) + apply (intro conjI, assumption+) + by (auto simp:set_ins_def) + show ?thesis + proof(rule reps_splitedI[where v = "(v - int (x + 1) - 1)"]) + show "(reps u (v - int (x + 1) - 1) xs \* zero (v - int (x + 1) - 1 + 1) \* + reps (v - int (x + 1) - 1 + 2) (1 + v) [Suc x]) s" + proof(rule sep_conjI) + from hhh(1) show "reps u (v - int (x + 1) - 1) xs s11" . + next + show "(zero (v - int (x + 1) - 1 + 1) \* reps (v - int (x + 1) - 1 + 2) (1 + v) [Suc x]) (s12 + (s13 + s2))" + proof(rule sep_conjI) + from hhh(2) show "zero (v - int (x + 1) - 1 + 1) s12" by smt + next + from hh(4) hhh(3) + show "reps (v - int (x + 1) - 1 + 2) (1 + v) [Suc x] (s13 + s2)" + apply (simp add:reps.simps del:ones_simps add:ones_rev) + by (smt hh(3) hh(4) hhh(4) hhh(5) hhh(6) sep_add_disjD sep_conjI sep_disj_addI1) + next + show "s12 ## s13 + s2" + by (metis hh(3) hhh(4) hhh(5) hhh(6) sep_add_commute sep_add_disjD + sep_add_disjI2 sep_disj_addD sep_disj_addD1 sep_disj_addI1 sep_disj_commuteI) + next + show "s12 + (s13 + s2) = s12 + (s13 + s2)" by metis + qed + next + show "s11 ## s12 + (s13 + s2)" + by (metis hh(3) hhh(4) hhh(5) hhh(6) sep_disj_addD1 sep_disj_addI1 sep_disj_addI3) + next + show "s = s11 + (s12 + (s13 + s2))" + by (smt hh(2) hh(3) hhh(4) hhh(5) hhh(6) sep_add_assoc sep_add_commute + sep_add_disjD sep_add_disjI2 sep_disj_addD1 sep_disj_addD2 + sep_disj_addI1 sep_disj_addI3 sep_disj_commuteI) + qed + next + from False show "xs \ []" . + next + show "[Suc x] \ []" by simp + qed +qed + +lemma cond_expand: "( \* p) s = (cond \ (p s))" + by (metis (full_types) condD pasrt_def sep_conj_commuteI + sep_conj_sep_emptyI sep_empty_def sep_globalise) + +lemma ones_rev1: + assumes "\ (1 + n) < m" + shows "(ones m n \* one (1 + n)) = (ones m (1 + n))" + by (insert ones_rev[OF assms, simplified], simp) + +lemma reps_one_abs: + assumes "(reps u v [k] \* one w) s" + "w = v + 1" + shows "(reps u w [Suc k]) s" + using assms unfolding assms + apply (simp add:reps.simps sep_conj_ac) + apply (subst (asm) sep_conj_cond)+ + apply (erule condE, simp) + by (simp add:ones_rev sep_conj_ac, sep_cancel+, smt) + +lemma reps'_reps_abs: + assumes "(reps' u v xs \* reps w x ys) s" + "w = v + 1" "ys \ []" + shows "(reps u x (xs@ys)) s" +proof(cases "xs = []") + case False + with assms + have h: "splited (xs@ys) xs ys" by (simp add:splited_def) + from assms(1)[unfolded assms(2)] + show ?thesis + apply (unfold reps_splited[OF h]) + apply (insert False, unfold reps'_def, simp) + apply (fwd reps_lenE, (subst (asm) sep_conj_cond)+) + by (erule condE, simp) +next + case True + with assms + show ?thesis + apply (simp add:reps'_def) + by (erule condE, simp) +qed + +lemma reps_one_abs1: + assumes "(reps u v (xs@[k]) \* one w) s" + "w = v + 1" + shows "(reps u w (xs@[Suc k])) s" +proof(cases "xs = []") + case True + with assms reps_one_abs + show ?thesis by simp +next + case False + hence "splited (xs@[k]) xs [k]" by (simp add:splited_def) + from assms(1)[unfolded reps_splited[OF this] assms(2)] + show ?thesis + apply (fwd reps_one_abs) + apply (fwd reps_reps'_abs) + apply (fwd reps'_reps_abs) + by (simp add:assms) +qed + +lemma tm_hoare_inc00: + assumes h: "a < length ks \ ks ! a = v" + shows "\st i \* ps u \* zero (u - 2) \* zero (u - 1) \* reps u ia ks \* fam_conj {ia<..} zero\ + i :[ Inc a ]: j + \st j \* + ps u \* + zero (u - 2) \* + zero (u - 1) \* + reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1) (ks[a := Suc v]) \* + fam_conj {(ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1)<..} zero\" + (is "\ ?P \ ?code \ ?Q \") +proof - + from h have "a < length ks" "ks ! a = v" by auto + from list_nth_expand[OF `a < length ks`] + have eq_ks: "ks = take a ks @ [ks!a] @ drop (Suc a) ks" . + from `a < length ks` have "ks \ []" by auto + hence "(reps u ia ks \* zero (ia + 1)) = reps' u (ia + 1) ks" + by (simp add:reps'_def) + also from eq_ks have "\ = reps' u (ia + 1) (take a ks @ [ks!a] @ drop (Suc a) ks)" by simp + also have "\ = (EXS m. reps' u (m - 1) (take a ks) \* + reps' m (ia + 1) (ks ! a # drop (Suc a) ks))" + by (simp add:reps'_append) + also have "\ = (EXS m. reps' u (m - 1) (take a ks) \* + reps' m (ia + 1) ([ks ! a] @ drop (Suc a) ks))" + by simp + also have "\ = (EXS m ma. reps' u (m - 1) (take a ks) \* + reps' m (ma - 1) [ks ! a] \* reps' ma (ia + 1) (drop (Suc a) ks))" + by (simp only:reps'_append sep_conj_exists) + finally have eq_s: "(reps u ia ks \* zero (ia + 1)) = \" . + { fix m ma + have eq_u: "-1 + u = u - 1" by smt + have " \st i \* + ps u \* + zero (u - 2) \* + zero (u - 1) \* + (reps' u (m - 1) (take a ks) \* + reps' m (ma - 1) [ks ! a] \* reps' ma (ia + 1) (drop (Suc a) ks)) \* + fam_conj {ia + 1<..} zero\ + i :[ Inc a ]: j + \st j \* + ps u \* + zero (u - 2) \* + zero (u - 1) \* + reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1) (ks[a := Suc v]) \* + fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1<..} zero\" + proof(cases "(drop (Suc a) ks) = []") + case True + { fix hc + have eq_1: "(1 + (m + int (ks ! a))) = (m + int (ks ! a) + 1)" by simp + have eq_2: "take a ks @ [Suc (ks ! a)] = ks[a := Suc v]" + apply (subst (3) eq_ks, unfold True, simp) + by (metis True append_Nil2 eq_ks h upd_conv_take_nth_drop) + assume h: "(fam_conj {1 + (m + int (ks ! a))<..} zero \* + reps u (1 + (m + int (ks ! a))) (take a ks @ [Suc (ks ! a)])) hc" + hence "(fam_conj {m + int (ks ! a) + 1<..} zero \* reps u (m + int (ks ! a) + 1) (ks[a := Suc v])) hc" + by (unfold eq_1 eq_2 , sep_cancel+) + } note eq_fam = this + show ?thesis + apply (unfold Inc_def, subst (3) reps'_def, simp add:True sep_conj_cond) + apply (intro tm.pre_condI, simp) + apply (subst fam_conj_interv_simp, simp, subst (3) zero_def) + apply hsteps + apply (subst reps'_sg, simp add:sep_conj_cond del:ones_simps) + apply (rule tm.pre_condI, simp del:ones_simps) + apply hsteps + apply (rule_tac p = " + st j' \* ps (1 + (m + int (ks ! a))) \* zero (u - 1) \* zero (u - 2) \* + reps u (1 + (m + int (ks ! a))) (take a ks @ [Suc (ks!a)]) + \* fam_conj {1 + (m + int (ks ! a))<..} zero + " in tm.pre_stren) + apply (hsteps) + apply (simp add:sep_conj_ac list_ext_lt[OF `a < length ks`], sep_cancel+) + apply (fwd eq_fam, sep_cancel+) + apply (simp del:ones_simps add:sep_conj_ac) + apply (sep_cancel+, prune) + apply ((fwd abs_reps')+, simp) + apply (fwd reps_one_abs abs_reps')+ + apply (subst (asm) reps'_def, simp) + by (subst fam_conj_interv_simp, simp add:sep_conj_ac) + next + case False + then obtain k' ks' where eq_drop: "drop (Suc a) ks = [k']@ks'" + by (metis append_Cons append_Nil list.exhaust) + from `a < length ks` + have eq_ks: "ks[a := Suc v] = (take a ks @ [Suc (ks ! a)]) @ (drop (Suc a) ks)" + apply (fold `ks!a = v`) + by (metis append_Cons append_Nil append_assoc upd_conv_take_nth_drop) + show ?thesis + apply (unfold Inc_def) + apply (unfold Inc_def eq_drop reps'_append, simp add:sep_conj_exists del:ones_simps) + apply (rule tm.precond_exI, subst (2) reps'_sg) + apply (subst sep_conj_cond)+ + apply (subst (1) ones_int_expand) + apply (rule tm.pre_condI, simp del:ones_simps) + apply hsteps + (* apply (hsteps hoare_locate_skip_gen[OF `a < length ks`]) *) + apply (subst reps'_sg, simp add:sep_conj_cond del:ones_simps) + apply (rule tm.pre_condI, simp del:ones_simps) + apply hsteps + apply (rule_tac p = "st j' \* + ps (2 + (m + int (ks ! a))) \* + reps (2 + (m + int (ks ! a))) ia (drop (Suc a) ks) \* zero (ia + 1) \* zero (ia + 2) \* + reps u (m + int (ks ! a)) (take a ks @ [ks!a]) \* zero (1 + (m + int (ks ! a))) \* + zero (u - 2) \* zero (u - 1) \* fam_conj {ia + 2<..} zero + " in tm.pre_stren) + apply (hsteps hoare_shift_right_cons_gen[OF False] + hoare_left_until_double_zero_gen[OF False]) + apply (rule_tac p = + "st j' \* ps (1 + (m + int (ks ! a))) \* + zero (u - 2) \* zero (u - 1) \* + reps u (1 + (m + int (ks ! a))) (take a ks @ [Suc (ks ! a)]) \* + zero (2 + (m + int (ks ! a))) \* + reps (3 + (m + int (ks ! a))) (ia + 1) (drop (Suc a) ks) \* fam_conj {ia + 1<..} zero + " in tm.pre_stren) + apply (hsteps) + apply (simp add:sep_conj_ac, sep_cancel+) + apply (unfold list_ext_lt[OF `a < length ks`], simp) + apply (fwd abs_reps')+ + apply(fwd reps'_reps_abs) + apply (subst eq_ks, simp) + apply (sep_cancel+) + apply (thin_tac "mb = 4 + (m + (int (ks ! a) + int k'))") + apply (thin_tac "ma = 2 + (m + int (ks ! a))", prune) + apply (simp add: int_add_ac sep_conj_ac, sep_cancel+) + apply (fwd reps_one_abs1, subst fam_conj_interv_simp, simp, sep_cancel+, smt) + apply (fwd abs_ones)+ + apply (fwd abs_reps') + apply (fwd abs_reps') + apply (fwd abs_reps') + apply (fwd abs_reps') + apply (unfold eq_drop, simp add:int_add_ac sep_conj_ac) + apply (sep_cancel+) + apply (fwd reps'_abs[where xs = "take a ks"]) + apply (fwd reps'_abs[where xs = "[k']"]) + apply (unfold reps'_def, simp add:int_add_ac sep_conj_ac) + apply (sep_cancel+) + by (subst (asm) fam_conj_interv_simp, simp, smt) + qed + } note h = this + show ?thesis + apply (subst fam_conj_interv_simp) + apply (rule_tac p = "st i \* ps u \* zero (u - 2) \* zero (u - 1) \* + (reps u ia ks \* zero (ia + 1)) \* fam_conj {ia + 1<..} zero" + in tm.pre_stren) + apply (unfold eq_s, simp only:sep_conj_exists) + apply (intro tm.precond_exI h) + by (sep_cancel+, unfold eq_s, simp) +qed + +declare ones_simps [simp del] + +lemma tm_hoare_inc01: + assumes "length ks \ a \ v = 0" + shows + "\st i \* ps u \* zero (u - 2) \* zero (u - 1) \* reps u ia ks \* fam_conj {ia<..} zero\ + i :[ Inc a ]: j + \st j \* + ps u \* + zero (u - 2) \* + zero (u - 1) \* + reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1) ((list_ext a ks)[a := Suc v]) \* + fam_conj {(ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1)<..} zero\" +proof - + from assms have "length ks \ a" "v = 0" by auto + show ?thesis + apply (rule_tac p = " + st i \* ps u \* zero (u - 2) \* zero (u - 1) \* (reps u ia ks \* <(ia = u + int (reps_len ks) - 1)>) \* + fam_conj {ia<..} zero + " in tm.pre_stren) + apply (subst sep_conj_cond)+ + apply (rule tm.pre_condI, simp) + apply (unfold Inc_def) + apply hstep + (* apply (hstep hoare_locate_set_gen[OF `length ks \ a`]) *) + apply (simp only:sep_conj_exists) + apply (intro tm.precond_exI) + my_block + fix m w + have "reps m w [list_ext a ks ! a] = + (ones m (m + int (list_ext a ks ! a)) \* <(w = m + int (list_ext a ks ! a))>)" + by (simp add:reps.simps) + my_block_end + apply (unfold this) + apply (subst sep_conj_cond)+ + apply (rule tm.pre_condI, simp) + apply (subst fam_conj_interv_simp) + apply (hsteps) + apply (subst fam_conj_interv_simp, simp) + apply (hsteps) + apply (rule_tac p = "st j' \* ps (m + int (list_ext a ks ! a) + 1) \* + zero (u - 2) \* zero (u - 1) \* + reps u (m + int (list_ext a ks ! a) + 1) + ((take a (list_ext a ks))@[Suc (list_ext a ks ! a)]) \* + fam_conj {(m + int (list_ext a ks ! a) + 1)<..} zero + " in tm.pre_stren) + apply (hsteps) + apply (simp add:sep_conj_ac, sep_cancel+) + apply (unfold `v = 0`, prune) + my_block + from `length ks \ a` have "list_ext a ks ! a = 0" + by (metis le_refl list_ext_tail) + from `length ks \ a` have "a < length (list_ext a ks)" + by (metis list_ext_len) + from reps_len_suc[OF this] + have eq_len: "int (reps_len (list_ext a ks)) = + int (reps_len (list_ext a ks[a := Suc (list_ext a ks ! a)])) - 1" + by smt + fix m w hc + assume h: "(fam_conj {m + int (list_ext a ks ! a) + 1<..} zero \* + reps u (m + int (list_ext a ks ! a) + 1) (take a (list_ext a ks) @ [Suc (list_ext a ks ! a)])) + hc" + then obtain s where + "(reps u (m + int (list_ext a ks ! a) + 1) (take a (list_ext a ks) @ [Suc (list_ext a ks ! a)])) s" + by (auto dest!:sep_conjD) + from reps_len_correct[OF this] + have "m = u + int (reps_len (take a (list_ext a ks) @ [Suc (list_ext a ks ! a)])) + - int (list_ext a ks ! a) - 2" by smt + from h [unfolded this] + have "(fam_conj {u + int (reps_len (list_ext a ks))<..} zero \* + reps u (u + int (reps_len (list_ext a ks))) (list_ext a ks[a := Suc 0])) + hc" + apply (unfold eq_len, simp) + my_block + from `a < length (list_ext a ks)` + have "take a (list_ext a ks) @ [Suc (list_ext a ks ! a)] = + list_ext a ks[a := Suc (list_ext a ks ! a)]" + by (smt `list_ext a ks ! a = 0` assms length_take list_ext_tail_expand list_update_length) + my_block_end + apply (unfold this) + my_block + have "-1 + (u + int (reps_len (list_ext a ks[a := Suc (list_ext a ks ! a)]))) = + u + (int (reps_len (list_ext a ks[a := Suc (list_ext a ks ! a)])) - 1)" by simp + my_block_end + apply (unfold this) + apply (sep_cancel+) + by (unfold `(list_ext a ks ! a) = 0`, simp) + my_block_end + apply (rule this, assumption) + apply (simp only:sep_conj_ac, sep_cancel+)+ + apply (fwd abs_reps')+ + apply (fwd reps_one_abs) + apply (fwd reps'_reps_abs) + apply (simp add:int_add_ac sep_conj_ac) + apply (sep_cancel+) + apply (subst fam_conj_interv_simp, simp add:sep_conj_ac, smt) + apply (fwd reps_lenE, (subst (asm) sep_conj_cond)+, erule condE, simp) + by (sep_cancel+) +qed + +lemma cond_eq: "(( \* p) s) = (b \ (p s))" +proof + assume "( \* p) s" + from condD[OF this] show " b \ p s" . +next + assume "b \ p s" + hence b and "p s" by auto + from `b` have "() 0" by (auto simp:pasrt_def) + moreover have "s = 0 + s" by auto + moreover have "0 ## s" by auto + moreover note `p s` + ultimately show "( \* p) s" by (auto intro!:sep_conjI) +qed + +lemma tm_hoare_dec_fail00: + assumes "a < length ks \ ks ! a = 0" + shows "\st i \* ps u \* zero (u - 2) \* zero (u - 1) \* reps u ia ks \* fam_conj {ia<..} zero\ + i :[ Dec a e ]: j + \st e \* ps u \* zero (u - 2) \* zero (u - 1) \* + reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks)) (list_ext a ks[a := 0]) \* + fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) <..} zero\" +proof - + from assms have "a < length ks" "ks!a = 0" by auto + from list_nth_expand[OF `a < length ks`] + have eq_ks: "ks = take a ks @ [ks ! a] @ drop (Suc a) ks" . + show ?thesis + proof(cases " drop (Suc a) ks = []") + case False + then obtain k' ks' where eq_drop: "drop (Suc a) ks = [k']@ks'" + by (metis append_Cons append_Nil list.exhaust) + show ?thesis + apply (unfold Dec_def, intro t_hoare_local) + apply (subst tassemble_to.simps(2), rule tm.code_exI, rule tm.code_extension) + apply (subst (1) eq_ks) + my_block + have "(reps u ia (take a ks @ [ks ! a] @ drop (Suc a) ks) \* fam_conj {ia<..} zero) = + (reps' u (ia + 1) ((take a ks @ [ks ! a]) @ drop (Suc a) ks) \* fam_conj {ia + 1<..} zero)" + apply (subst fam_conj_interv_simp) + by (unfold reps'_def, simp add:sep_conj_ac) + my_block_end + apply (unfold this) + apply (subst reps'_append) + apply (unfold eq_drop) + apply (subst (2) reps'_append) + apply (simp only:sep_conj_exists, intro tm.precond_exI) + apply (subst (2) reps'_def, simp add:reps.simps ones_simps) + apply (subst reps'_append, simp only:sep_conj_exists, intro tm.precond_exI) + apply hstep + (* apply (hstep hoare_locate_skip_gen[OF `a < length ks`]) *) + my_block + fix m mb + have "(reps' mb (m - 1) [ks ! a]) = (reps mb (m - 2) [ks!a] \* zero (m - 1))" + by (simp add:reps'_def, smt) + my_block_end + apply (unfold this) + apply hstep + (* apply (hstep hoare_if_reps_nz_false_gen[OF `ks!a = 0`]) *) + apply (simp only:reps.simps(2), simp add:`ks!a = 0`) + apply (rule_tac p = "st j'b \* + ps mb \* + reps u mb ((take a ks)@[ks ! a]) \* <(m - 2 = mb)> \* + zero (m - 1) \* + zero (u - 1) \* + one m \* + zero (u - 2) \* + ones (m + 1) (m + int k') \* + <(-2 + ma = m + int k')> \* zero (ma - 1) \* reps' ma (ia + 1) ks' \* fam_conj {ia + 1<..} zero" + in tm.pre_stren) + apply hsteps + apply (simp add:sep_conj_ac, sep_cancel+) + apply (subgoal_tac "m + int k' = ma - 2", simp) + apply (fwd abs_ones)+ + apply (subst (asm) sep_conj_cond)+ + apply (erule condE, auto) + apply (fwd abs_reps')+ + apply (subgoal_tac "ma = m + int k' + 2", simp) + apply (fwd abs_reps')+ + my_block + from `a < length ks` + have "list_ext a ks = ks" by (auto simp:list_ext_def) + my_block_end + apply (simp add:this) + apply (subst eq_ks, simp add:eq_drop `ks!a = 0`) + apply (subst (asm) reps'_def, simp) + apply (subst fam_conj_interv_simp, simp add:sep_conj_ac, sep_cancel+) + apply (metis append_Cons assms eq_Nil_appendI eq_drop eq_ks list_update_id) + apply (clarsimp) + apply (subst (asm) sep_conj_cond)+ + apply (erule condE, clarsimp) + apply (subst (asm) sep_conj_cond)+ + apply (erule condE, clarsimp) + apply (simp add:sep_conj_ac, sep_cancel+) + apply (fwd abs_reps')+ + by (fwd reps'_reps_abs, simp add:`ks!a = 0`) + next + case True + show ?thesis + apply (unfold Dec_def, intro t_hoare_local) + apply (subst tassemble_to.simps(2), rule tm.code_exI, rule tm.code_extension) + apply (subst (1) eq_ks, unfold True, simp) + my_block + have "(reps u ia (take a ks @ [ks ! a]) \* fam_conj {ia<..} zero) = + (reps' u (ia + 1) (take a ks @ [ks ! a]) \* fam_conj {ia + 1<..} zero)" + apply (unfold reps'_def, subst fam_conj_interv_simp) + by (simp add:sep_conj_ac) + my_block_end + apply (subst (1) this) + apply (subst reps'_append) + apply (simp only:sep_conj_exists, intro tm.precond_exI) + apply (subst fam_conj_interv_simp, simp) + my_block + have "(zero (2 + ia)) = (tm (2 + ia) Bk)" + by (simp add:zero_def) + my_block_end my_note eq_z = this + apply hstep + (* apply (hstep hoare_locate_skip_gen[OF `a < length ks`]) *) + my_block + fix m + have "(reps' m (ia + 1) [ks ! a]) = (reps m ia [ks!a] \* zero (ia + 1))" + by (simp add:reps'_def) + my_block_end + apply (unfold this, prune) + apply hstep + (* apply (hstep hoare_if_reps_nz_false_gen[OF `ks!a = 0`]) *) + apply (simp only:reps.simps(2), simp add:`ks!a = 0`) + apply (rule_tac p = "st j'b \* ps m \* (reps u m ((take a ks)@[ks!a]) \* <(ia = m)>) + \* zero (ia + 1) \* zero (u - 1) \* + zero (2 + ia) \* zero (u - 2) \* fam_conj {2 + ia<..} zero" + in tm.pre_stren) + apply hsteps + apply (simp add:sep_conj_ac) + apply ((subst (asm) sep_conj_cond)+, erule condE, simp) + my_block + from `a < length ks` have "list_ext a ks = ks" by (metis list_ext_lt) + my_block_end + apply (unfold this, simp) + apply (subst fam_conj_interv_simp) + apply (subst fam_conj_interv_simp, simp) + apply (simp only:sep_conj_ac, sep_cancel+) + apply (subst eq_ks, unfold True `ks!a = 0`, simp) + apply (metis True append_Nil2 assms eq_ks list_update_same_conv) + apply (simp add:sep_conj_ac) + apply (subst (asm) sep_conj_cond)+ + apply (erule condE, simp, thin_tac "ia = m") + apply (fwd abs_reps')+ + apply (simp add:sep_conj_ac int_add_ac, sep_cancel+) + apply (unfold reps'_def, simp) + by (metis sep.mult_commute) + qed +qed + +lemma tm_hoare_dec_fail01: + assumes "length ks \ a" + shows "\st i \* ps u \* zero (u - 2) \* zero (u - 1) \* reps u ia ks \* fam_conj {ia<..} zero\ + i :[ Dec a e ]: j + \st e \* ps u \* zero (u - 2) \* zero (u - 1) \* + reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks)) (list_ext a ks[a := 0]) \* + fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) <..} zero\" + apply (unfold Dec_def, intro t_hoare_local) + apply (subst tassemble_to.simps(2), rule tm.code_exI, rule tm.code_extension) + apply (rule_tac p = "st i \* ps u \* zero (u - 2) \* + zero (u - 1) \* reps u ia ks \* fam_conj {ia<..} zero \* + <(ia = u + int (reps_len ks) - 1)>" in tm.pre_stren) + apply hstep + (* apply (hstep hoare_locate_set_gen[OF `length ks \ a`]) *) + apply (simp only:sep_conj_exists, intro tm.precond_exI) + my_block + from assms + have "list_ext a ks ! a = 0" by (metis le_refl list_ext_tail) + my_block_end my_note is_z = this + apply (subst fam_conj_interv_simp) + apply hstep + (* apply (hstep hoare_if_reps_nz_false_gen[OF is_z]) *) + apply (unfold is_z) + apply (subst (1) reps.simps) + apply (rule_tac p = "st j'b \* ps m \* reps u m (take a (list_ext a ks) @ [0]) \* zero (w + 1) \* + <(w = m + int 0)> \* zero (u - 1) \* + fam_conj {w + 1<..} zero \* zero (u - 2) \* + <(ia = u + int (reps_len ks) - 1)>" in tm.pre_stren) + my_block + have "(take a (list_ext a ks)) @ [0] \ []" by simp + my_block_end + apply hsteps + (* apply (hsteps hoare_left_until_double_zero_gen[OF this]) *) + apply (simp add:sep_conj_ac) + apply prune + apply (subst (asm) sep_conj_cond)+ + apply (elim condE, simp add:sep_conj_ac, prune) + my_block + fix m w ha + assume h1: "w = m \ ia = u + int (reps_len ks) - 1" + and h: "(ps u \* + st e \* + zero (u - 1) \* + zero (m + 1) \* + fam_conj {m + 1<..} zero \* zero (u - 2) \* reps u m (take a (list_ext a ks) @ [0])) ha" + from h1 have eq_w: "w = m" and eq_ia: "ia = u + int (reps_len ks) - 1" by auto + from h obtain s' where "reps u m (take a (list_ext a ks) @ [0]) s'" + by (auto dest!:sep_conjD) + from reps_len_correct[OF this] + have eq_m: "m = u + int (reps_len (take a (list_ext a ks) @ [0])) - 1" . + from h[unfolded eq_m, simplified] + have "(ps u \* + st e \* + zero (u - 1) \* + zero (u - 2) \* + fam_conj {u + (-1 + int (reps_len (list_ext a ks)))<..} zero \* + reps u (u + (-1 + int (reps_len (list_ext a ks)))) (list_ext a ks[a := 0])) ha" + apply (sep_cancel+) + apply (subst fam_conj_interv_simp, simp) + my_block + from `length ks \ a` have "list_ext a ks[a := 0] = list_ext a ks" + by (metis is_z list_update_id) + my_block_end + apply (unfold this) + my_block + from `length ks \ a` is_z + have "take a (list_ext a ks) @ [0] = list_ext a ks" + by (metis list_ext_tail_expand) + my_block_end + apply (unfold this) + by (simp add:sep_conj_ac, sep_cancel+, smt) + my_block_end + apply (rule this, assumption) + apply (sep_cancel+)[1] + apply (subst (asm) sep_conj_cond)+ + apply (erule condE, prune, simp) + my_block + fix s m + assume "(reps' u (m - 1) (take a (list_ext a ks)) \* ones m m \* zero (m + 1)) s" + hence "reps' u (m + 1) (take a (list_ext a ks) @ [0]) s" + apply (unfold reps'_append) + apply (rule_tac x = m in EXS_intro) + by (subst (2) reps'_def, simp add:reps.simps) + my_block_end + apply (rotate_tac 1, fwd this) + apply (subst (asm) reps'_def, simp add:sep_conj_ac) + my_block + fix s + assume h: "(st i \* ps u \* zero (u - 2) \* zero (u - 1) \* + reps u ia ks \* fam_conj {ia<..} zero) s" + then obtain s' where "reps u ia ks s'" by (auto dest!:sep_conjD) + from reps_len_correct[OF this] have eq_ia: "ia = u + int (reps_len ks) - 1" . + from h + have "(st i \* ps u \* zero (u - 2) \* zero (u - 1) \* reps u ia ks \* + fam_conj {ia<..} zero \* <(ia = u + int (reps_len ks) - 1)>) s" + by (unfold eq_ia, simp) + my_block_end + by (rule this, assumption) + +lemma t_hoare_label1: + "(\l. l = i \ \st l \* p\ l :[ c l ]: j \st k \* q\) \ + \st l \* p \ + i:[(TLabel l; c l)]:j + \st k \* q\" +by (unfold tassemble_to.simps, intro tm.code_exI tm.code_condI, clarify, auto) + +lemma tm_hoare_dec_fail1: + assumes "a < length ks \ ks ! a = 0 \ length ks \ a" + shows "\st i \* ps u \* zero (u - 2) \* zero (u - 1) \* reps u ia ks \* fam_conj {ia<..} zero\ + i :[ Dec a e ]: j + \st e \* ps u \* zero (u - 2) \* zero (u - 1) \* + reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks)) (list_ext a ks[a := 0]) \* + fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) <..} zero\" + using assms +proof + assume "a < length ks \ ks ! a = 0" + thus ?thesis + by (rule tm_hoare_dec_fail00) +next + assume "length ks \ a" + thus ?thesis + by (rule tm_hoare_dec_fail01) +qed + +lemma shift_left_nil_gen[step]: + assumes "u = v" + shows "\st i \* ps u \* zero v\ + i :[shift_left]:j + \st j \* ps u \* zero v\" + apply(unfold assms shift_left_def, intro t_hoare_local t_hoare_label, clarify, + rule t_hoare_label_last, simp, clarify, prune, simp) + by hstep + +lemma tm_hoare_dec_suc1: + assumes "a < length ks \ ks ! a = Suc v" + shows "\st i \* ps u \* zero (u - 2) \* zero (u - 1) \* reps u ia ks \* fam_conj {ia<..} zero\ + i :[ Dec a e ]: j + \st j \* ps u \* zero (u - 2) \* zero (u - 1) \* + reps u (ia - 1) (list_ext a ks[a := v]) \* + fam_conj {ia - 1<..} zero\" +proof - + from assms have "a < length ks" " ks ! a = Suc v" by auto + from list_nth_expand[OF `a < length ks`] + have eq_ks: "ks = take a ks @ [ks ! a] @ drop (Suc a) ks" . + show ?thesis + proof(cases " drop (Suc a) ks = []") + case False + then obtain k' ks' where eq_drop: "drop (Suc a) ks = [k']@ks'" + by (metis append_Cons append_Nil list.exhaust) + show ?thesis + apply (unfold Dec_def, intro t_hoare_local) + apply (subst tassemble_to.simps(2), rule tm.code_exI) + apply (subst (1) eq_ks) + my_block + have "(reps u ia (take a ks @ [ks ! a] @ drop (Suc a) ks) \* fam_conj {ia<..} zero) = + (reps' u (ia + 1) ((take a ks @ [ks ! a]) @ drop (Suc a) ks) \* fam_conj {ia + 1<..} zero)" + apply (subst fam_conj_interv_simp) + by (unfold reps'_def, simp add:sep_conj_ac) + my_block_end + apply (unfold this) + apply (subst reps'_append) + apply (unfold eq_drop) + apply (subst (2) reps'_append) + apply (simp only:sep_conj_exists, intro tm.precond_exI) + apply (subst (2) reps'_def, simp add:reps.simps ones_simps) + apply (subst reps'_append, simp only:sep_conj_exists, intro tm.precond_exI) + apply (rule_tac q = + "st l \* + ps mb \* + zero (u - 1) \* + reps' u (mb - 1) (take a ks) \* + reps' mb (m - 1) [ks ! a] \* + one m \* + zero (u - 2) \* + ones (m + 1) (m + int k') \* + <(-2 + ma = m + int k')> \* zero (ma - 1) \* reps' ma (ia + 1) ks' \* fam_conj {ia + 1<..} zero" + in tm.sequencing) + apply (rule tm.code_extension) + apply hstep + (* apply (hstep hoare_locate_skip_gen[OF `a < length ks`]) *) + apply (subst (2) reps'_def, simp) + my_block + fix i j l m mb + from `ks!a = (Suc v)` have "ks!a \ 0" by simp + from hoare_if_reps_nz_true[OF this, where u = mb and v = "m - 2"] + have "\st i \* ps mb \* reps mb (-2 + m) [ks ! a]\ + i :[ if_reps_nz l ]: j + \st l \* ps mb \* reps mb (-2 + m) [ks ! a]\" + by smt + my_block_end + apply (hgoto this) + apply (simp add:sep_conj_ac, sep_cancel+) + apply (subst reps'_def, simp add:sep_conj_ac) + apply (rule tm.code_extension1) + apply (rule t_hoare_label1, simp, prune) + apply (subst (2) reps'_def, simp add:reps.simps) + apply (rule_tac p = "st j' \* ps mb \* zero (u - 1) \* reps' u (mb - 1) (take a ks) \* + ((ones mb (mb + int (ks ! a)) \* <(-2 + m = mb + int (ks ! a))>) \* zero (mb + int (ks ! a) + 1)) \* + one (mb + int (ks ! a) + 2) \* zero (u - 2) \* + ones (mb + int (ks ! a) + 3) (mb + int (ks ! a) + int k' + 2) \* + <(-2 + ma = m + int k')> \* zero (ma - 1) \* reps' ma (ia + 1) ks' \* fam_conj {ia + 1<..} zero + " in tm.pre_stren) + apply hsteps + (* apply (simp add:sep_conj_ac) *) + apply (unfold `ks!a = Suc v`) + my_block + fix mb + have "(ones mb (mb + int (Suc v))) = (ones mb (mb + int v) \* one (mb + int (Suc v)))" + by (simp add:ones_rev) + my_block_end + apply (unfold this, prune) + apply hsteps + apply (rule_tac p = "st j'a \* + ps (mb + int (Suc v) + 2) \* zero (mb + int (Suc v) + 1) \* + reps (mb + int (Suc v) + 2) ia (drop (Suc a) ks) \* zero (ia + 1) \* zero (ia + 2) \* + zero (mb + int (Suc v)) \* + ones mb (mb + int v) \* + zero (u - 1) \* + reps' u (mb - 1) (take a ks) \* + zero (u - 2) \* fam_conj {ia + 2<..} zero + " in tm.pre_stren) + apply hsteps + (* apply (hsteps hoare_shift_left_cons_gen[OF False]) *) + apply (rule_tac p = "st j'a \* ps (ia - 1) \* zero (u - 2) \* zero (u - 1) \* + reps u (ia - 1) (take a ks @ [v] @ drop (Suc a) ks) \* + zero ia \* zero (ia + 1) \* zero (ia + 2) \* + fam_conj {ia + 2<..} zero + " in tm.pre_stren) + apply hsteps + apply (simp add:sep_conj_ac) + apply (subst fam_conj_interv_simp) + apply (subst fam_conj_interv_simp) + apply (subst fam_conj_interv_simp) + apply (simp add:sep_conj_ac) + apply (sep_cancel+) + my_block + have "take a ks @ v # drop (Suc a) ks = list_ext a ks[a := v]" + proof - + from `a < length ks` have "list_ext a ks = ks" by (metis list_ext_lt) + hence "list_ext a ks[a:=v] = ks[a:=v]" by simp + moreover from `a < length ks` have "ks[a:=v] = take a ks @ v # drop (Suc a) ks" + by (metis upd_conv_take_nth_drop) + ultimately show ?thesis by metis + qed + my_block_end + apply (unfold this, sep_cancel+, smt) + apply (simp add:sep_conj_ac) + apply (fwd abs_reps')+ + apply (simp add:sep_conj_ac int_add_ac) + apply (sep_cancel+) + apply (subst (asm) reps'_def, simp add:sep_conj_ac) + apply (subst (asm) sep_conj_cond)+ + apply (erule condE, clarsimp) + apply (simp add:sep_conj_ac, sep_cancel+) + apply (fwd abs_ones)+ + apply (fwd abs_reps')+ + apply (subst (asm) reps'_def, simp) + apply (subst (asm) fam_conj_interv_simp) + apply (simp add:sep_conj_ac int_add_ac eq_drop reps'_def) + apply (subst (asm) sep_conj_cond)+ + apply (erule condE, clarsimp) + by (simp add:sep_conj_ac int_add_ac) + next + case True + show ?thesis + apply (unfold Dec_def, intro t_hoare_local) + apply (subst tassemble_to.simps(2), rule tm.code_exI) + apply (subst (1) eq_ks, simp add:True) + my_block + have "(reps u ia (take a ks @ [ks ! a]) \* fam_conj {ia<..} zero) = + (reps' u (ia + 1) (take a ks @ [ks ! a]) \* fam_conj {ia + 1<..} zero)" + apply (subst fam_conj_interv_simp) + by (unfold reps'_def, simp add:sep_conj_ac) + my_block_end + apply (unfold this) + apply (subst reps'_append) + apply (simp only:sep_conj_exists, intro tm.precond_exI) + apply (rule_tac q = "st l \* ps m \* zero (u - 1) \* reps' u (m - 1) (take a ks) \* + reps' m (ia + 1) [ks ! a] \* zero (2 + ia) \* zero (u - 2) \* fam_conj {2 + ia<..} zero" + in tm.sequencing) + apply (rule tm.code_extension) + apply (subst fam_conj_interv_simp, simp) + apply hsteps + (* apply (hstep hoare_locate_skip_gen[OF `a < length ks`]) *) + my_block + fix m + have "(reps' m (ia + 1) [ks ! a]) = + (reps m ia [ks!a] \* zero (ia + 1))" + by (unfold reps'_def, simp) + my_block_end + apply (unfold this) + my_block + fix i j l m + from `ks!a = (Suc v)` have "ks!a \ 0" by simp + my_block_end + apply (hgoto hoare_if_reps_nz_true_gen) + apply (rule tm.code_extension1) + apply (rule t_hoare_label1, simp) + apply (thin_tac "la = j'", prune) + apply (subst (1) reps.simps) + apply (subst sep_conj_cond)+ + apply (rule tm.pre_condI, simp) + apply hsteps + apply (unfold `ks!a = Suc v`) + my_block + fix m + have "(ones m (m + int (Suc v))) = (ones m (m + int v) \* one (m + int (Suc v)))" + by (simp add:ones_rev) + my_block_end + apply (unfold this) + apply hsteps + apply (rule_tac p = "st j'a \* ps (m + int v) \* zero (u - 2) \* zero (u - 1) \* + reps u (m + int v) (take a ks @ [v]) \* zero (m + (1 + int v)) \* + zero (2 + (m + int v)) \* zero (3 + (m + int v)) \* + fam_conj {3 + (m + int v)<..} zero + " in tm.pre_stren) + apply hsteps + apply (simp add:sep_conj_ac, sep_cancel+) + my_block + have "take a ks @ [v] = list_ext a ks[a := v]" + proof - + from True `a < length ks` have "ks = take a ks @ [ks!a]" + by (metis append_Nil2 eq_ks) + hence "ks[a:=v] = take a ks @ [v]" + by (metis True `a < length ks` upd_conv_take_nth_drop) + moreover from `a < length ks` have "list_ext a ks = ks" + by (metis list_ext_lt) + ultimately show ?thesis by simp + qed + my_block_end my_note eq_l = this + apply (unfold this) + apply (subst fam_conj_interv_simp) + apply (subst fam_conj_interv_simp) + apply (subst fam_conj_interv_simp) + apply (simp add:sep_conj_ac, sep_cancel, smt) + apply (simp add:sep_conj_ac int_add_ac)+ + apply (sep_cancel+) + apply (fwd abs_reps')+ + apply (fwd reps'_reps_abs) + by (simp add:eq_l) + qed +qed + +lemma hoare_cfill_until_one: + "\st i \* ps v \* one (u - 1) \* zeros u v\ + i :[ cfill_until_one ]: j + \st j \* ps (u - 1) \* ones (u - 1) v \" +proof(induct u v rule:zeros_rev_induct) + case (Base x y) + thus ?case + apply (subst sep_conj_cond)+ + apply (rule tm.pre_condI, simp add:ones_simps) + apply (unfold cfill_until_one_def, intro t_hoare_local t_hoare_label, rule t_hoare_label_last, simp+) + by hstep +next + case (Step x y) + show ?case + apply (rule_tac q = "st i \* ps (y - 1) \* one (x - 1) \* zeros x (y - 1) \* one y" in tm.sequencing) + apply (subst cfill_until_one_def, intro t_hoare_local t_hoare_label, rule t_hoare_label_last, simp+) + apply hsteps + my_block + fix i j l + have "\st i \* ps (y - 1) \* one (x - 1) \* zeros x (y - 1)\ + i :[ jmp l ]: j + \st l \* ps (y - 1) \* one (x - 1) \* zeros x (y - 1)\" + apply (case_tac "(y - 1) < x", simp add:zeros_simps) + apply (subst sep_conj_cond)+ + apply (rule tm.pre_condI, simp) + apply hstep + apply (drule_tac zeros_rev, simp) + by hstep + my_block_end + apply (hstep this) + (* The next half *) + apply (hstep Step(2), simp add:sep_conj_ac, sep_cancel+) + by (insert Step(1), simp add:ones_rev sep_conj_ac) +qed + +declare zeros.simps [simp del] zeros_simps[simp del] + +lemma hoare_cmove: + assumes "w \ k" + shows "\st i \* ps (v + 2 + int w) \* zero (u - 1) \* + reps u (v - int w) [k - w] \* zeros (v - int w + 1) (v + 1) \* + one (v + 2) \* ones (v + 3) (v + 2 + int w) \* zeros (v + 3 + int w) (v + int(reps_len [k]) + 1)\ + i :[cmove]: j + \st j \* ps (u - 1) \* zero (u - 1) \* reps u u [0] \* zeros (u + 1) (v + 1) \* + reps (v + 2) (v + int (reps_len [k]) + 1) [k]\" + using assms +proof(induct "k - w" arbitrary: w) + case (0 w) + hence "w = k" by auto + show ?case + apply (simp add: `w = k` del:zeros.simps zeros_simps) + apply (unfold cmove_def, intro t_hoare_local t_hoare_label, rule t_hoare_label_last, simp+) + apply (simp add:reps_len_def reps_sep_len_def reps_ctnt_len_def del:zeros_simps zeros.simps) + apply (rule_tac p = "st i \* ps (v + 2 + int k) \* zero (u - 1) \* + reps u u [0] \* zeros (u + 1) (v + 1) \* + ones (v + 2) (v + 2 + int k) \* zeros (v + 3 + int k) (2 + (v + int k)) \* + <(u = v - int k)>" + in tm.pre_stren) + my_block + fix i j + have "\st i \* ps (v + 2 + int k) \* zeros (u + 1) (v + 1) \* ones (v + 2) (v + 2 + int k) + \* <(u = v - int k)>\ + i :[ left_until_zero ]: j + \st j \* ps (v + 1) \* zeros (u + 1) (v + 1) \* ones (v + 2) (v + 2 + int k) + \* <(u = v - int k)>\" + apply (subst sep_conj_cond)+ + apply (rule tm.pre_condI, simp) + my_block + have "(zeros (v - int k + 1) (v + 1)) = (zeros (v - int k + 1) v \* zero (v + 1))" + by (simp only:zeros_rev, smt) + my_block_end + apply (unfold this) + by hsteps + my_block_end + apply (hstep this) + my_block + fix i j + have "\st i \* ps (v + 1) \* reps u u [0] \* zeros (u + 1) (v + 1)\ + i :[left_until_one]:j + \st j \* ps u \* reps u u [0] \* zeros (u + 1) (v + 1)\" + apply (simp add:reps.simps ones_simps) + by hsteps + my_block_end + apply (hsteps this) + apply ((subst (asm) sep_conj_cond)+, erule condE, clarsimp) + apply (fwd abs_reps')+ + apply (simp only:sep_conj_ac int_add_ac, sep_cancel+) + apply (simp add:int_add_ac sep_conj_ac zeros_simps) + apply (simp add:sep_conj_ac int_add_ac, sep_cancel+) + apply (fwd reps_lenE) + apply (subst (asm) sep_conj_cond)+ + apply (erule condE, clarsimp) + apply (subgoal_tac "v = u + int k + int (reps_len [0]) - 1", clarsimp) + apply (simp add:reps_len_sg) + apply (fwd abs_ones)+ + apply (fwd abs_reps')+ + apply (simp add:sep_conj_ac int_add_ac) + apply (sep_cancel+) + apply (simp add:reps.simps, smt) + by (clarsimp) +next + case (Suc k' w) + from `Suc k' = k - w` `w \ k` + have h: "k' = k - (Suc w)" "Suc w \ k" by auto + show ?case + apply (rule tm.sequencing[OF _ Suc(1)[OF h(1, 2)]]) + apply (unfold cmove_def, intro t_hoare_local t_hoare_label, rule t_hoare_label_last, simp+) + apply (simp add:reps_len_def reps_sep_len_def reps_ctnt_len_def del:zeros_simps zeros.simps) + my_block + fix i j + have "\st i \* ps (v + 2 + int w) \* zeros (v - int w + 1) (v + 1) \* + one (v + 2) \* ones (v + 3) (v + 2 + int w) \ + i :[left_until_zero]: j + \st j \* ps (v + 1) \* zeros (v - int w + 1) (v + 1) \* + one (v + 2) \* ones (v + 3) (v + 2 + int w) \" + my_block + have "(one (v + 2) \* ones (v + 3) (v + 2 + int w)) = + ones (v + 2) (v + 2 + int w)" + by (simp only:ones_simps, smt) + my_block_end + apply (unfold this) + my_block + have "(zeros (v - int w + 1) (v + 1)) = (zeros (v - int w + 1) v \* zero (v + 1))" + by (simp only:zeros_rev, simp) + my_block_end + apply (unfold this) + by hsteps + my_block_end + apply (hstep this) + my_block + fix i j + have "\st i \* ps (v + 1) \* reps u (v - int w) [k - w] \* zeros (v - int w + 1) (v + 1)\ + i :[left_until_one]: j + \st j \* ps (v - int w) \* reps u (v - int w) [k - w] \* zeros (v - int w + 1) (v + 1)\" + apply (simp add:reps.simps ones_rev) + apply (subst sep_conj_cond)+ + apply (rule tm.pre_condI, clarsimp) + apply (subgoal_tac "u + int (k - w) = v - int w", simp) + defer + apply simp + by hsteps + my_block_end + apply (hstep this) + my_block + have "(reps u (v - int w) [k - w]) = (reps u (v - (1 + int w)) [k - Suc w] \* one (v - int w))" + apply (subst (1 2) reps.simps) + apply (subst sep_conj_cond)+ + my_block + have "((v - int w = u + int (k - w))) = + (v - (1 + int w) = u + int (k - Suc w))" + apply auto + apply (smt Suc.prems h(2)) + by (smt Suc.prems h(2)) + my_block_end + apply (simp add:this) + my_block + fix b p q + assume "(b \ (p::tassert) = q)" + have "( \* p) = ( \* q)" + by (metis `b \ p = q` cond_eqI) + my_block_end + apply (rule this) + my_block + assume "v - (1 + int w) = u + int (k - Suc w)" + hence "v = 1 + int w + u + int (k - Suc w)" by auto + my_block_end + apply (simp add:this) + my_block + have "\ (u + int (k - w)) < u" by auto + my_block_end + apply (unfold ones_rev[OF this]) + my_block + from Suc (2, 3) have "(u + int (k - w) - 1) = (u + int (k - Suc w))" + by auto + my_block_end + apply (unfold this) + my_block + from Suc (2, 3) have "(u + int (k - w)) = (1 + (u + int (k - Suc w)))" + by auto + my_block_end + by (unfold this, simp) + my_block_end + apply (unfold this) + my_block + fix i j + have "\st i \* ps (v - int w) \* + (reps u (v - (1 + int w)) [k - Suc w] \* one (v - int w))\ + i :[ move_left]: j + \st j \* ps (v - (1 + int w)) \* + (reps u (v - (1 + int w)) [k - Suc w] \* one (v - int w))\" + apply (simp add:reps.simps ones_rev) + apply (subst sep_conj_cond)+ + apply (rule tm.pre_condI, clarsimp) + apply (subgoal_tac " u + int (k - Suc w) = v - (1 + int w)", simp) + defer + apply simp + apply hsteps + by (simp add:sep_conj_ac, sep_cancel+, smt) + my_block_end + apply (hstep this) + my_block + fix i' j' + have "\st i' \* ps (v - (1 + int w)) \* reps u (v - (1 + int w)) [k - Suc w]\ + i' :[ if_zero j ]: j' + \st j' \* ps (v - (1 + int w)) \* reps u (v - (1 + int w)) [k - Suc w]\" + apply (simp add:reps.simps ones_rev) + apply (subst sep_conj_cond)+ + apply (rule tm.pre_condI, clarsimp) + apply (subgoal_tac " u + int (k - Suc w) = v - (1 + int w)", simp) + defer + apply simp + by hstep + my_block_end + apply (hstep this) + my_block + fix i j + have "\st i \* ps (v - (1 + int w)) \* reps u (v - (1 + int w)) [k - Suc w]\ + i :[ move_right ]: j + \st j \* ps (v - int w) \* reps u (v - (1 + int w)) [k - Suc w] \" + apply (simp add:reps.simps ones_rev) + apply (subst sep_conj_cond)+ + apply (rule tm.pre_condI, clarsimp) + apply (subgoal_tac " u + int (k - Suc w) = v - (1 + int w)", simp) + defer + apply simp + by hstep + my_block_end + apply (hsteps this) + my_block + fix i j + have "\st i \* ps (v - int w) \* one (v + 2) \* + zero (v - int w) \* zeros (v - int w + 1) (v + 1)\ + i :[right_until_one]: j + \st j \* ps (v + 2) \* one (v + 2) \* zero (v - int w) \* zeros (v - int w + 1) (v + 1)\" + my_block + have "(zero (v - int w) \* zeros (v - int w + 1) (v + 1)) = + (zeros (v - int w) (v + 1))" + by (simp add:zeros_simps) + my_block_end + apply (unfold this) + by hsteps + my_block_end + apply (hstep this) + my_block + from Suc(2, 3) have "w < k" by auto + hence "(zeros (v + 3 + int w) (2 + (v + int k))) = + (zero (v + 3 + int w) \* zeros (4 + (v + int w)) (2 + (v + int k)))" + by (simp add:zeros_simps) + my_block_end + apply (unfold this) + my_block + fix i j + have "\st i \* ps (v + 2) \* zero (v + 3 + int w) \* zeros (4 + (v + int w)) (2 + (v + int k)) \* + one (v + 2) \* ones (v + 3) (v + 2 + int w)\ + i :[right_until_zero]: j + \st j \* ps (v + 3 + int w) \* zero (v + 3 + int w) \* zeros (4 + (v + int w)) (2 + (v + int k)) \* + one (v + 2) \* ones (v + 3) (v + 2 + int w)\" + my_block + have "(one (v + 2) \* ones (v + 3) (v + 2 + int w)) = + (ones (v + 2) (v + 2 + int w))" + by (simp add:ones_simps, smt) + my_block_end + apply (unfold this) + by hsteps + my_block_end + apply (hsteps this, simp only:sep_conj_ac) + apply (sep_cancel+, simp add:sep_conj_ac) + my_block + fix s + assume "(zero (v - int w) \* zeros (v - int w + 1) (v + 1)) s" + hence "zeros (v - int w) (v + 1) s" + by (simp add:zeros_simps) + my_block_end + apply (fwd this) + my_block + fix s + assume "(one (v + 3 + int w) \* ones (v + 3) (v + 2 + int w)) s" + hence "ones (v + 3) (3 + (v + int w)) s" + by (simp add:ones_rev sep_conj_ac, smt) + my_block_end + apply (fwd this) + by (simp add:sep_conj_ac, smt) +qed + +lemma hoare_copy: + shows + "\st i \* ps u \* zero (u - 2) \* zero (u - 1) \* reps u v [k] \* zero (v + 1) \* + zeros (v + 2) (v + int(reps_len [k]) + 1)\ + i :[copy]: j + \st j \* ps u \* zero (u - 2) \* zero (u - 1) \* reps u v [k] \* zero (v + 1) \* + reps (v + 2) (v + int (reps_len [k]) + 1) [k]\" + apply (unfold copy_def) + my_block + fix i j + have + "\st i \* ps u \* reps u v [k] \* zero (v + 1) \* zeros (v + 2) (v + int(reps_len [k]) + 1)\ + i :[cinit]: j + \st j \* ps (v + 2) \* reps u v [k] \* zero (v + 1) \* + one (v + 2) \* zeros (v + 3) (v + int(reps_len [k]) + 1)\" + apply (unfold cinit_def) + apply (simp add:reps.simps) + apply (subst sep_conj_cond)+ + apply (rule tm.pre_condI, simp) + apply hsteps + apply (simp add:sep_conj_ac) + my_block + have "(zeros (u + int k + 2) (u + int k + int (reps_len [k]) + 1)) = + (zero (u + int k + 2) \* zeros (u + int k + 3) (u + int k + int (reps_len [k]) + 1))" + by (smt reps_len_sg zeros_step_simp) + my_block_end + apply (unfold this) + apply hstep + by (simp add:sep_conj_ac, sep_cancel+, smt) + my_block_end + apply (hstep this) + apply (rule_tac p = "st j' \* ps (v + 2) \* reps u v [k] \* zero (v + 1) \* + one (v + 2) \* zeros (v + 3) (v + int (reps_len [k]) + 1) \* zero (u - 2) \* zero (u - 1) \* + <(v = u + int (reps_len [k]) - 1)> + " in tm.pre_stren) + my_block + fix i j + from hoare_cmove[where w = 0 and k = k and i = i and j = j and v = v and u = u] + have "\st i \* ps (v + 2) \* zero (u - 1) \* reps u v [k] \* zero (v + 1) \* + one (v + 2) \* zeros (v + 3) (v + int(reps_len [k]) + 1)\ + i :[cmove]: j + \st j \* ps (u - 1) \* zero (u - 1) \* reps u u [0] \* zeros (u + 1) (v + 1) \* + reps (v + 2) (v + int (reps_len [k]) + 1) [k]\" + by (auto simp:ones_simps zeros_simps) + my_block_end + apply (hstep this) + apply (hstep, simp) + my_block + have "reps u u [0] = one u" by (simp add:reps.simps ones_simps) + my_block_end my_note eq_repsz = this + apply (unfold this) + apply (hstep) + apply (subst reps.simps, simp add: ones_simps) + apply hsteps + apply (subst sep_conj_cond)+ + apply (rule tm.pre_condI, simp del:zeros.simps zeros_simps) + apply (thin_tac "int (reps_len [k]) = 1 + int k \ v = u + int (reps_len [k]) - 1") + my_block + have "(zeros (u + 1) (u + int k + 1)) = (zeros (u + 1) (u + int k) \* zero (u + int k + 1))" + by (simp only:zeros_rev, smt) + my_block_end + apply (unfold this) + apply (hstep, simp) + my_block + fix i j + from hoare_cfill_until_one[where v = "u + int k" and u = "u + 1"] + have "\st i \* ps (u + int k) \* one u \* zeros (u + 1) (u + int k)\ + i :[ cfill_until_one ]: j + \st j \* ps u \* ones u (u + int k) \" + by simp + my_block_end + apply (hstep this, simp add:sep_conj_ac reps.simps ones_simps) + apply (simp add:sep_conj_ac reps.simps ones_simps) + apply (subst sep_conj_cond)+ + apply (subst (asm) sep_conj_cond)+ + apply (rule condI) + apply (erule condE, simp) + apply (simp add: reps_len_def reps_sep_len_def reps_ctnt_len_def) + apply (sep_cancel+) + by (erule condE, simp) + +end + diff -r 77daf1b85cf0 -r a5f5b9336007 thys2/Hoare_tm_basis.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/Hoare_tm_basis.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,2857 @@ +header {* + Separation logic for TM +*} + +theory Hoare_tm_basis +imports Hoare_gen My_block Data_slot MLs Term_pat (* BaseSS *) Subgoal Sort_ops + Thm_inst +begin + +section {* Aux lemmas on seperation algebra *} + +lemma cond_eq: "(( \* p) s) = (b \ (p s))" +proof + assume "( \* p) s" + from condD[OF this] show " b \ p s" . +next + assume "b \ p s" + hence b and "p s" by auto + from `b` have "() 0" by (auto simp:pasrt_def) + moreover have "s = 0 + s" by auto + moreover have "0 ## s" by auto + moreover note `p s` + ultimately show "( \* p) s" by (auto intro!:sep_conjI) +qed + +lemma cond_eqI: + assumes h: "b \ r = s" + shows "( ** r) = ( ** s)" +proof(cases b) + case True + from h[OF this] show ?thesis by simp +next + case False + thus ?thesis + by (unfold sep_conj_def set_ins_def pasrt_def, auto) +qed + +lemma EXS_intro: + assumes h: "(P x) s" + shows "((EXS x. P(x))) s" + by (smt h pred_ex_def sep_conj_impl) + +lemma EXS_elim: + assumes "(EXS x. P x) s" + obtains x where "P x s" + by (metis assms pred_ex_def) + +lemma EXS_eq: + fixes x + assumes h: "Q (p x)" + and h1: "\ y s. \p y s\ \ y = x" + shows "p x = (EXS x. p x)" +proof + fix s + show "p x s = (EXS x. p x) s" + proof + assume "p x s" + thus "(EXS x. p x) s" by (auto simp:pred_ex_def) + next + assume "(EXS x. p x) s" + thus "p x s" + proof(rule EXS_elim) + fix y + assume "p y s" + from this[unfolded h1[OF this]] show "p x s" . + qed + qed +qed + +section {* The TM assembly language *} + +subsection {* The TM assembly language *} + +datatype taction = W0 | W1 | L | R + +datatype tstate = St nat + +fun nat_of :: "tstate \ nat" +where "nat_of (St n) = n" + +declare [[coercion_enabled]] + +declare [[coercion "St :: nat \ tstate"]] + +type_synonym tm_inst = "(taction \ tstate) \ (taction \ tstate)" + +datatype Block = Oc | Bk + +datatype tpg = + TInstr tm_inst + | TLabel tstate + | TSeq tpg tpg + | TLocal "(tstate \ tpg)" + +notation TLocal (binder "TL " 10) + +abbreviation tprog_instr :: "tm_inst \ tpg" ("\ _" [61] 61) +where "\ i \ TInstr i" + +abbreviation tprog_seq :: "tpg \ tpg \ tpg" (infixr ";" 52) +where "c1 ; c2 \ TSeq c1 c2" + +subsection {* The notion of assembling *} + +datatype tresource = + TM int Block + | TC nat tm_inst + | TAt nat + | TPos int + | TFaults nat + +type_synonym tassert = "tresource set \ bool" + +definition "sg e = (\ s. s = e)" + +primrec tassemble_to :: "tpg \ nat \ nat \ tassert" + where + "tassemble_to (TInstr ai) i j = (sg ({TC i ai}) ** <(j = i + 1)>)" | + "tassemble_to (TSeq p1 p2) i j = (EXS j'. (tassemble_to p1 i j') ** (tassemble_to p2 j' j))" | + "tassemble_to (TLocal fp) i j = (EXS l. (tassemble_to (fp l) i j))" | + "tassemble_to (TLabel l) i j = <(i = j \ j = nat_of l)>" + +declare tassemble_to.simps [simp del] + +lemmas tasmp = tassemble_to.simps (2, 3, 4) + +abbreviation tasmb_to :: "nat \ tpg \ nat \ tassert" ("_ :[ _ ]: _" [60, 60, 60] 60) + where "i :[ tpg ]: j \ tassemble_to tpg i j" + +section {* Automatic checking of assemblility *} + +subsection {* Basic theories *} + +text {* @{text cpg} is the type for skeleton assembly language. Every constructor + corresponds to one in the definition of @{text tpg} *} + +datatype cpg = + CInstr tm_inst + | CLabel nat + | CSeq cpg cpg + | CLocal cpg + +text {* Conversion from @{text cpg} to @{text tpg}*} + +fun c2t :: "tstate list \ cpg \ tpg" where + "c2t env (CInstr ((act0, St st0), (act1, St st1))) = + TInstr ((act0, env!st0), (act1, env!st1))" | + "c2t env (CLabel l) = TLabel (env!l)" | + "c2t env (CSeq cpg1 cpg2) = TSeq (c2t env cpg1) (c2t env cpg2)" | + "c2t env (CLocal cpg) = TLocal (\ x. c2t (x#env) cpg)" + +text {* Well formedness checking of @{text cpg} *} + +datatype status = Free | Bound + +text {* @{text wf_cpg_test} is the checking function *} + +fun wf_cpg_test :: "status list \ cpg \ (bool \ status list)" where + "wf_cpg_test sts (CInstr ((a0, l0), (a1, l1))) = ((l0 < length sts \ l1 < length sts), sts)" | + "wf_cpg_test sts (CLabel l) = ((l < length sts) \ sts!l = Free, sts[l:=Bound])" | + "wf_cpg_test sts (CSeq c1 c2) = (let (b1, sts1) = wf_cpg_test sts c1; + (b2, sts2) = wf_cpg_test sts1 c2 in + (b1 \ b2, sts2))" | + "wf_cpg_test sts (CLocal body) = (let (b, sts') = (wf_cpg_test (Free#sts) body) in + (b, tl sts'))" + +text {* + The meaning the following @{text "c2p"} has to be understood together with + the following lemma @{text "wf_cpg_test_correct"}. The intended use of @{text "c2p"} + is when the elements of @{text "sts"} are all @{text "Free"}, in which case, + the precondition on @{text "f"}, i.e. + @{text "\ x. ((length x = length sts \ + (\ k < length sts. sts!k = Bound \ (x!k = f i k))"} + is trivially true. +*} +definition + "c2p sts i cpg j = + (\ f. \ x. ((length x = length sts \ + (\ k < length sts. sts!k = Bound \ (x!k = f i k))) + \ (\ s. (i:[(c2t x cpg)]:j) s)))" + +instantiation status :: order +begin + definition less_eq_status_def: "((st1::status) \ st2) = (st1 = Free \ st2 = Bound)" + definition less_status_def: "((st1::status) < st2) = (st1 \ st2 \ st1 \ st2)" +instance +proof + fix x y + show "(x < (y::status)) = (x \ y \ \ y \ x)" + by (metis less_eq_status_def less_status_def status.distinct(1)) +next + fix x show "x \ (x::status)" + by (metis (full_types) less_eq_status_def status.exhaust) +next + fix x y z + assume "x \ y" "y \ (z::status)" show "x \ (z::status)" + by (metis `x \ y` `y \ z` less_eq_status_def status.distinct(1)) +next + fix x y + assume "x \ y" "y \ (x::status)" show "x = y" + by (metis `x \ y` `y \ x` less_eq_status_def status.distinct(1)) +qed +end + +instantiation list :: (order)order +begin + definition "((sts::('a::order) list) \ sts') = + ((length sts = length sts') \ (\ i < length sts. sts!i \ sts'!i))" + definition "((sts::('a::order) list) < sts') = ((sts \ sts') \ sts \ sts')" + + lemma anti_sym: assumes h: "x \ (y::'a list)" "y \ x" + shows "x = y" + proof - + from h have "length x = length y" + by (metis less_eq_list_def) + moreover from h have " (\ i < length x. x!i = y!i)" + by (metis (full_types) antisym_conv less_eq_list_def) + ultimately show ?thesis + by (metis nth_equalityI) + qed + + lemma refl: "x \ (x::('a::order) list)" + apply (unfold less_eq_list_def) + by (metis order_refl) + + instance + proof + fix x y + show "((x::('a::order) list) < y) = (x \ y \ \ y \ x)" + proof + assume h: "x \ y \ \ y \ x" + have "x \ y" + proof + assume "x = y" with h have "\ (x \ x)" by simp + with refl show False by auto + qed + moreover from h have "x \ y" by blast + ultimately show "x < y" by (auto simp:less_list_def) + next + assume h: "x < y" + hence hh: "x \ y" + by (metis less_list_def) + moreover have "\ y \ x" + proof + assume "y \ x" + from anti_sym[OF hh this] have "x = y" . + with h show False + by (metis less_list_def) + qed + ultimately show "x \ y \ \ y \ x" by auto + qed + next + fix x from refl show "(x::'a list) \ x" . + next + fix x y assume "(x::'a list) \ y" "y \ x" + from anti_sym[OF this] show "x = y" . + next + fix x y z + assume h: "(x::'a list) \ y" "y \ z" + show "x \ z" + proof - + from h have "length x = length z" by (metis less_eq_list_def) + moreover from h have "\ i < length x. x!i \ z!i" + by (metis less_eq_list_def order_trans) + ultimately show "x \ z" + by (metis less_eq_list_def) + qed + qed +end + +lemma sts_bound_le: "sts \ sts[l := Bound]" +proof - + have "length sts = length (sts[l := Bound])" + by (metis length_list_update) + moreover have "\ i < length sts. sts!i \ (sts[l:=Bound])!i" + proof - + { fix i + assume "i < length sts" + have "sts ! i \ sts[l := Bound] ! i" + proof(cases "l < length sts") + case True + note le_l = this + show ?thesis + proof(cases "l = i") + case True with le_l + have "sts[l := Bound] ! i = Bound" by auto + thus ?thesis by (metis less_eq_status_def) + next + case False + with le_l have "sts[l := Bound] ! i = sts!i" by auto + thus ?thesis by auto + qed + next + case False + hence "sts[l := Bound] = sts" by auto + thus ?thesis by auto + qed + } thus ?thesis by auto + qed + ultimately show ?thesis by (metis less_eq_list_def) +qed + +lemma sts_tl_le: + assumes "sts \ sts'" + shows "tl sts \ tl sts'" +proof - + from assms have "length (tl sts) = length (tl sts')" + by (metis (hide_lams, no_types) length_tl less_eq_list_def) + moreover from assms have "\ i < length (tl sts). (tl sts)!i \ (tl sts')!i" + by (smt calculation length_tl less_eq_list_def nth_tl) + ultimately show ?thesis + by (metis less_eq_list_def) +qed + +lemma wf_cpg_test_le: + assumes "wf_cpg_test sts cpg = (True, sts')" + shows "sts \ sts'" using assms +proof(induct cpg arbitrary:sts sts') + case (CInstr instr sts sts') + obtain a0 l0 a1 l1 where eq_instr: "instr = ((a0, St l0), (a1, St l1))" + by (metis surj_pair tstate.exhaust) + from CInstr[unfolded this] + show ?case by simp +next + case (CLabel l sts sts') + thus ?case by (auto simp:sts_bound_le) +next + case (CLocal body sts sts') + from this(2) + obtain sts1 where h: "wf_cpg_test (Free # sts) body = (True, sts1)" "sts' = tl sts1" + by (auto split:prod.splits) + from CLocal(1)[OF this(1)] have "Free # sts \ sts1" . + from sts_tl_le[OF this] + have "sts \ tl sts1" by simp + from this[folded h(2)] + show ?case . +next + case (CSeq cpg1 cpg2 sts sts') + from this(3) + show ?case + by (auto split:prod.splits dest!:CSeq(1, 2)) +qed + +lemma c2p_assert: + assumes "(c2p [] i cpg j)" + shows "\ s. (i :[(c2t [] cpg)]: j) s" +proof - + from assms obtain f where + h [rule_format]: + "\ x. length x = length [] \ (\k (x ! k = f i k)) \ + (\ s. (i :[ c2t [] cpg ]: j) s)" + by (unfold c2p_def, auto) + have "length [] = length [] \ (\k ([] ! k = f i k))" + by auto + from h[OF this] obtain s where "(i :[ c2t [] cpg ]: j) s" by blast + thus ?thesis by auto +qed + +definition "sts_disj sts sts' = ((length sts = length sts') \ + (\ i < length sts. \(sts!i = Bound \ sts'!i = Bound)))" + +instantiation list :: (plus)plus +begin + fun plus_list :: "'a list \ 'a list \ 'a list" where + "plus_list [] ys = []" | + "plus_list (x#xs) [] = []" | + "plus_list (x#xs) (y#ys) = ((x + y)#plus_list xs ys)" + instance .. +end + +instantiation list :: (minus)minus +begin + fun minus_list :: "'a list \ 'a list \ 'a list" where + "minus_list [] ys = []" | + "minus_list (x#xs) [] = []" | + "minus_list (x#xs) (y#ys) = ((x - y)#minus_list xs ys)" + instance .. +end + +instantiation status :: minus +begin + fun minus_status :: "status \ status \ status" where + "minus_status Bound Bound = Free" | + "minus_status Bound Free = Bound" | + "minus_status Free x = Free " + instance .. +end + +instantiation status :: plus +begin + fun plus_status :: "status \ status \ status" where + "plus_status Free x = x" | + "plus_status Bound x = Bound" + instance .. +end + +lemma length_sts_plus: + assumes "length (sts1 :: status list) = length sts2" + shows "length (sts1 + sts2) = length sts1" + using assms +proof(induct sts1 arbitrary: sts2) + case Nil + thus ?case + by (metis plus_list.simps(1)) +next + case (Cons s' sts' sts2) + thus ?case + proof(cases "sts2 = []") + case True + with Cons show ?thesis by auto + next + case False + then obtain s'' sts'' where eq_sts2: "sts2 = s''#sts''" + by (metis neq_Nil_conv) + with Cons + show ?thesis + by (metis length_Suc_conv list.inject plus_list.simps(3)) + qed +qed + +lemma nth_sts_plus: + assumes "i < length ((sts1::status list) + sts2)" + shows "(sts1 + sts2)!i = sts1!i + sts2!i" + using assms +proof(induct sts1 arbitrary:i sts2) + case (Nil i sts2) + thus ?case by auto +next + case (Cons s' sts' i sts2) + show ?case + proof(cases "sts2 = []") + case True + with Cons show ?thesis by auto + next + case False + then obtain s'' sts'' where eq_sts2: "sts2 = s''#sts''" + by (metis neq_Nil_conv) + with Cons + show ?thesis + by (smt list.size(4) nth_Cons' plus_list.simps(3)) + qed +qed + +lemma nth_sts_minus: + assumes "i < length ((sts1::status list) - sts2)" + shows "(sts1 - sts2)!i = sts1!i - sts2!i" + using assms +proof(induct arbitrary:i rule:minus_list.induct) + case (3 x xs y ys i) + show ?case + proof(cases i) + case 0 + thus ?thesis by simp + next + case (Suc k) + with 3(2) have "k < length (xs - ys)" by auto + from 3(1)[OF this] and Suc + show ?thesis + by auto + qed +qed auto + +fun taddr :: "tresource \ nat" where + "taddr (TC i instr) = i" + +lemma tassemble_to_range: + assumes "(i :[tpg]: j) s" + shows "(i \ j) \ (\ r \ s. i \ taddr r \ taddr r < j)" + using assms +proof(induct tpg arbitrary: i j s) + case (TInstr instr i j s) + obtain a0 l0 a1 l1 where "instr = ((a0, l0), (a1, l1))" + by (metis pair_collapse) + with TInstr + show ?case + apply (simp add:tassemble_to.simps sg_def) + by (smt `instr = ((a0, l0), a1, l1)` cond_eq cond_true_eq1 + empty_iff insert_iff le_refl lessI sep.mult_commute taddr.simps) +next + case (TLabel l i j s) + thus ?case + apply (simp add:tassemble_to.simps) + by (smt equals0D pasrt_def set_zero_def) +next + case (TSeq c1 c2 i j s) + from TSeq(3) obtain s1 s2 j' where + h: "(i :[ c1 ]: j') s1" "(j' :[ c2 ]: j) s2" "s1 ## s2" "s = s1 + s2" + by (auto simp:tassemble_to.simps elim!:EXS_elim sep_conjE) + show ?case + proof - + { fix r + assume "r \ s" + with h (3, 4) have "r \ s1 \ r \ s2" + by (auto simp:set_ins_def) + hence "i \ j \ i \ taddr r \ taddr r < j" + proof + assume " r \ s1" + from TSeq(1)[OF h(1)] + have "i \ j'" "(\r\s1. i \ taddr r \ taddr r < j')" by auto + from this(2)[rule_format, OF `r \ s1`] + have "i \ taddr r" "taddr r < j'" by auto + with TSeq(2)[OF h(2)] + show ?thesis by auto + next + assume "r \ s2" + from TSeq(2)[OF h(2)] + have "j' \ j" "(\r\s2. j' \ taddr r \ taddr r < j)" by auto + from this(2)[rule_format, OF `r \ s2`] + have "j' \ taddr r" "taddr r < j" by auto + with TSeq(1)[OF h(1)] + show ?thesis by auto + qed + } thus ?thesis + by (smt TSeq.hyps(1) TSeq.hyps(2) h(1) h(2)) + qed +next + case (TLocal body i j s) + from this(2) obtain l s' where "(i :[ body l ]: j) s" + by (simp add:tassemble_to.simps, auto elim!:EXS_elim) + from TLocal(1)[OF this] + show ?case by auto +qed + +lemma c2p_seq: + assumes "c2p sts1 i cpg1 j'" + "c2p sts2 j' cpg2 j" + "sts_disj sts1 sts2" + shows "(c2p (sts1 + sts2) i (CSeq cpg1 cpg2) j)" +proof - + from assms(1)[unfolded c2p_def] + obtain f1 where + h1[rule_format]: + "\x. length x = length sts1 \ (\k (x ! k = f1 i k)) \ + Ex (i :[ c2t x cpg1 ]: j')" by blast + from assms(2)[unfolded c2p_def] + obtain f2 where h2[rule_format]: + "\x. length x = length sts2 \ (\k (x ! k = f2 j' k)) \ + Ex (j' :[ c2t x cpg2 ]: j)" by blast + from assms(3)[unfolded sts_disj_def] + have h3: "length sts1 = length sts2" + and h4[rule_format]: + "(\i (sts1 ! i = Bound \ sts2 ! i = Bound))" by auto + let ?f = "\ i k. if (sts1!k = Bound) then f1 i k else f2 j' k" + { fix x + assume h5: "length x = length (sts1 + sts2)" and + h6[rule_format]: "(\k x ! k = ?f i k)" + obtain s1 where h_s1: "(i :[ c2t x cpg1 ]: j') s1" + proof(atomize_elim, rule h1) + from h3 h5 have "length x = length sts1" + by (metis length_sts_plus) + moreover { + fix k assume hh: "k + (\k (x ! k = f1 i k))" + by blast + qed + obtain s2 where h_s2: "(j' :[ c2t x cpg2 ]: j) s2" + proof(atomize_elim, rule h2) + from h3 h5 have "length x = length sts2" + by (metis length_sts_plus) + moreover { + fix k + assume hh: "k + (\i (sts1 ! i = Bound \ sts2 ! i = Bound))` + calculation nth_sts_plus plus_status.simps(1) status.distinct(1) status.exhaust) + from h6[OF p1 p2] have "x ! k = (if sts1 ! k = Bound then f1 i k else f2 j' k)" . + moreover from h4[OF hh(1)[folded h3]] hh(2) have "sts1!k \ Bound" by auto + ultimately have "x!k = f2 j' k" by simp + } ultimately show "length x = length sts2 \ + (\k (x ! k = f2 j' k))" + by blast + qed + have h_s12: "s1 ## s2" + proof - + { fix r assume h: "r \ s1" "r \ s2" + with h_s1 h_s2 + have "False"by (smt tassemble_to_range) + } thus ?thesis by (auto simp:set_ins_def) + qed + have "(EXS j'. i :[ c2t x cpg1 ]: j' \* j' :[ c2t x cpg2 ]: j) (s1 + s2)" + proof(rule_tac x = j' in EXS_intro) + from h_s1 h_s2 h_s12 + show "(i :[ c2t x cpg1 ]: j' \* j' :[ c2t x cpg2 ]: j) (s1 + s2)" + by (metis sep_conjI) + qed + hence "\ s. (i :[ c2t x (CSeq cpg1 cpg2) ]: j) s" + by (auto simp:tassemble_to.simps) + } + hence "\f. \x. length x = length (sts1 + sts2) \ + (\k x ! k = f i k) \ + Ex (i :[ c2t x (CSeq cpg1 cpg2) ]: j)" + by (rule_tac x = ?f in exI, auto) + thus ?thesis + by(unfold c2p_def, auto) +qed + +lemma plus_list_len: + "length ((sts1::status list) + sts2) = min (length sts1) (length sts2)" + by(induct rule:plus_list.induct, auto) + +lemma minus_list_len: + "length ((sts1::status list) - sts2) = min (length sts1) (length sts2)" + by(induct rule:minus_list.induct, auto) + +lemma sts_le_comb: + assumes "sts1 \ sts2" + and "sts2 \ sts3" + shows "sts_disj (sts2 - sts1) (sts3 - sts2)" and + "(sts3 - sts1) = (sts2 - sts1) + (sts3 - sts2)" +proof - + from assms + have h1: "length sts1 = length sts2" "\i sts2 ! i" + and h2: "length sts2 = length sts3" "\i sts3 ! i" + by (unfold less_eq_list_def, auto) + have "sts_disj (sts2 - sts1) (sts3 - sts2)" + proof - + from h1(1) h2(1) + have "length (sts2 - sts1) = length (sts3 - sts2)" + by (metis minus_list_len) + moreover { + fix i + assume lt_i: "i sts2 ! i" "sts2 ! i \ sts3 ! i" . + moreover have "(sts2 - sts1) ! i = sts2!i - sts1!i" + by (metis lt_i nth_sts_minus) + moreover have "(sts3 - sts2)!i = sts3!i - sts2!i" + by (metis `length (sts2 - sts1) = length (sts3 - sts2)` lt_i nth_sts_minus) + ultimately have " \ ((sts2 - sts1) ! i = Bound \ (sts3 - sts2) ! i = Bound)" + apply (cases "sts2!i", cases "sts1!i", cases "sts3!i", simp, simp) + apply (cases "sts3!i", simp, simp) + apply (cases "sts1!i", cases "sts3!i", simp, simp) + by (cases "sts3!i", simp, simp) + } ultimately show ?thesis by (unfold sts_disj_def, auto) + qed + moreover have "(sts3 - sts1) = (sts2 - sts1) + (sts3 - sts2)" + proof(induct rule:nth_equalityI) + case 1 + show ?case by (metis h1(1) h2(1) length_sts_plus minus_list_len) + next + case 2 + { fix i + assume lt_i: "i i. \ j. (c2p (sts' - sts) i cpg j))" + using assms +proof(induct cpg arbitrary:sts sts') + case (CInstr instr sts sts') + obtain a0 l0 a1 l1 where eq_instr: "instr = ((a0, St l0), (a1, St l1))" + by (metis surj_pair tstate.exhaust) + show ?case + proof(unfold eq_instr c2p_def, clarsimp simp:tassemble_to.simps) + fix i + let ?a = "(Suc i)" and ?f = "\ i k. St i" + show "\a f. \x. length x = length (sts' - sts) \ + (\k x ! k = f i k) \ + Ex (sg {TC i ((a0, (x ! l0)), a1, (x ! l1))} \* <(a = (Suc i))>)" + proof(rule_tac x = ?a in exI, rule_tac x = ?f in exI, default, clarsimp) + fix x + let ?j = "Suc i" + let ?s = " {TC i ((a0, x ! l0), a1, x ! l1)}" + have "(sg {TC i ((a0, x ! l0), a1, x ! l1)} \* <(Suc i = Suc i)>) ?s" + by (simp add:tassemble_to.simps sg_def) + thus "Ex (sg {TC i ((a0, (x ! l0)), a1, (x ! l1))})" + by auto + qed + qed +next + case (CLabel l sts sts') + show ?case + proof + fix i + from CLabel + have h1: "l < length sts" + and h2: "sts ! l = Free" + and h3: "sts[l := Bound] = sts'" by auto + let ?f = "\ i k. St i" + have "\a f. \x. length x = length (sts' - sts) \ + (\k x ! k = f (i::nat) k) \ + Ex (<(i = a \ a = nat_of (x ! l))>)" + proof(rule_tac x = i in exI, rule_tac x = ?f in exI, clarsimp) + fix x + assume h[rule_format]: + "\k x ! k = St i" + from h1 h3 have p1: "l < length (sts' - sts)" + by (metis length_list_update min_max.inf.idem minus_list_len) + from p1 h2 h3 have p2: "(sts' - sts)!l = Bound" + by (metis h1 minus_status.simps(2) nth_list_update_eq nth_sts_minus) + from h[OF p1 p2] have "(<(i = nat_of (x ! l))>) 0" + by (simp add:set_ins_def) + thus "\ s. (<(i = nat_of (x ! l))>) s" by auto + qed + thus "\a. c2p (sts' - sts) i (CLabel l) a" + by (auto simp:c2p_def tassemble_to.simps) + qed +next + case (CSeq cpg1 cpg2 sts sts') + show ?case + proof + fix i + from CSeq(3)[unfolded wf_cpg_test.simps] + obtain b1 sts1 + where LetE: "(let (b2, y) = wf_cpg_test sts1 cpg2 in (b1 \ b2, y)) = (True, sts')" + "(b1, sts1) = wf_cpg_test sts cpg1" + by (auto simp:Let_def split:prod.splits) + show "\ j. c2p (sts' - sts) i (CSeq cpg1 cpg2) j" + proof - + from LetE(1) + obtain b2 where h: "(b2, sts') = wf_cpg_test sts1 cpg2" "b1=True" "b2=True" + by (atomize_elim, unfold Let_def, auto split:prod.splits) + from wf_cpg_test_le[OF LetE(2)[symmetric, unfolded h(2)]] + have sts_le1: "sts \ sts1" . + from CSeq(1)[OF LetE(2)[unfolded h(2), symmetric], rule_format, of i] + obtain j1 where h1: "(c2p (sts1 - sts) i cpg1 j1)" by blast + from wf_cpg_test_le[OF h(1)[symmetric, unfolded h(3)]] + have sts_le2: "sts1 \ sts'" . + from sts_le_comb[OF sts_le1 sts_le2] + have hh: "sts_disj (sts1 - sts) (sts' - sts1)" + "sts' - sts = (sts1 - sts) + (sts' - sts1)" . + from CSeq(2)[OF h(1)[symmetric, unfolded h(3)], rule_format, of j1] + obtain j2 where h2: "(c2p (sts' - sts1) j1 cpg2 j2)" by blast + have "c2p (sts' - sts) i (CSeq cpg1 cpg2) j2" + by(unfold hh(2), rule c2p_seq[OF h1 h2 hh(1)]) + thus ?thesis by blast + qed + qed +next + case (CLocal body sts sts') + from this(2) obtain b sts1 s where + h: "wf_cpg_test (Free # sts) body = (True, sts1)" + "sts' = tl sts1" + by (unfold wf_cpg_test.simps, auto split:prod.splits) + from wf_cpg_test_le[OF h(1), unfolded less_eq_list_def] h(2) + obtain s where eq_sts1: "sts1 = s#sts'" + by (metis Suc_length_conv list.size(4) tl.simps(2)) + from CLocal(1)[OF h(1)] have p1: "\i. \a. c2p (sts1 - (Free # sts)) i body a" . + show ?case + proof + fix i + from p1[rule_format, of i] obtain j where "(c2p (sts1 - (Free # sts)) i body) j" by blast + then obtain f where hh [rule_format]: + "\x. length x = length (sts1 - (Free # sts)) \ + (\k x ! k = f i k) \ + (\s. (i :[ c2t x body ]: j) s)" + by (auto simp:c2p_def) + let ?f = "\ i k. f i (Suc k)" + have "\j f. \x. length x = length (sts' - sts) \ + (\k x ! k = f i k) \ + (\s. (i :[ (TL xa. c2t (xa # x) body) ]: j) s)" + proof(rule_tac x = j in exI, rule_tac x = ?f in exI, default, clarsimp) + fix x + assume h1: "length x = length (sts' - sts)" + and h2: "\k x ! k = f i (Suc k)" + let ?l = "f i 0" let ?x = " ?l#x" + from wf_cpg_test_le[OF h(1)] have "length (Free#sts) = length sts1" + by (unfold less_eq_list_def, simp) + with h1 h(2) have q1: "length (?l # x) = length (sts1 - (Free # sts))" + by (smt Suc_length_conv length_Suc_conv list.inject list.size(4) + minus_list.simps(3) minus_list_len tl.simps(2)) + have q2: "(\k (f i 0 # x) ! k = f i k)" + proof - + { fix k + assume lt_k: "ks. (i :[ (TL xa. c2t (xa # x) body) ]: j) s)" + apply (simp add:tassemble_to.simps) + by (rule_tac x = s in exI, rule_tac x = "?l::tstate" in EXS_intro, simp) + qed + thus "\j. c2p (sts' - sts) i (CLocal body) j" + by (auto simp:c2p_def) + qed +qed + +lemma + assumes "wf_cpg_test [] cpg = (True, sts')" + and "tpg = c2t [] cpg" + shows "\ i. \ j s. ((i:[tpg]:j) s)" +proof + fix i + have eq_sts_minus: "(sts' - []) = []" + by (metis list.exhaust minus_list.simps(1) minus_list.simps(2)) + from wf_cpg_test_correct[OF assms(1), rule_format, of i] + obtain j where "c2p (sts' - []) i cpg j" by auto + from c2p_assert [OF this[unfolded eq_sts_minus]] + obtain s where "(i :[ c2t [] cpg ]: j) s" by blast + from this[folded assms(2)] + show " \ j s. ((i:[tpg]:j) s)" by blast +qed + +lemma replicate_nth: "(replicate n x @ sts) ! (l + n) = sts!l" + by (smt length_replicate nth_append) + +lemma replicate_update: + "(replicate n x @ sts)[l + n := v] = replicate n x @ sts[l := v]" + by (smt length_replicate list_update_append) + +lemma nth_app: "length xs \ a \ (xs @ ys)!a = ys!(a - length xs)" + by (metis not_less nth_append) + +lemma l_n_v_orig: + assumes "l0 < length env" + and "t \ l0" + shows "(take t env @ es @ drop t env) ! (l0 + length es) = env ! l0" +proof - + from assms(1, 2) have "t < length env" by auto + hence h: "env = take t env @ drop t env" + "length (take t env) = t" + apply (metis append_take_drop_id) + by (smt `t < length env` length_take) + with assms(2) have eq_sts_l: "env!l0 = (drop t env)!(l0 - t)" + by (metis nth_app) + from h(2) have "length (take t env @ es) = t + length es" + by (metis length_append length_replicate nat_add_commute) + moreover from assms(2) have "t + (length es) \ l0 + (length es)" by auto + ultimately have "((take t env @ es) @ drop t env)!(l0 + length es) = + (drop t env)!(l0+ length es - (t + length es))" + by (smt length_replicate length_splice nth_append) + with eq_sts_l[symmetric, unfolded assms] + show ?thesis by auto +qed + +lemma l_n_v: + assumes "l < length sts" + and "sts!l = v" + and "t \ l" + shows "(take t sts @ replicate n x @ drop t sts) ! (l + n) = v" +proof - + from l_n_v_orig[OF assms(1) assms(3), of "replicate n x"] + and assms(2) + show ?thesis by auto +qed + +lemma l_n_v_s: + assumes "l < length sts" + and "t \ l" + shows "(take t sts @ sts0 @ drop t sts)[l + length sts0 := v] = + take t (sts[l:=v])@ sts0 @ drop t (sts[l:=v])" +proof - + let ?n = "length sts0" + from assms(1, 2) have "t < length sts" by auto + hence h: "sts = take t sts @ drop t sts" + "length (take t sts) = t" + apply (metis append_take_drop_id) + by (smt `t < length sts` length_take) + with assms(2) have eq_sts_l: "sts[l:=v] = take t sts @ drop t sts [(l - t) := v]" + by (smt list_update_append) + with h(2) have eq_take_drop_t: "take t (sts[l:=v]) = take t sts" + "drop t (sts[l:=v]) = (drop t sts)[l - t:=v]" + apply (metis append_eq_conv_conj) + by (metis append_eq_conv_conj eq_sts_l h(2)) + from h(2) have "length (take t sts @ sts0) = t + ?n" + by (metis length_append length_replicate nat_add_commute) + moreover from assms(2) have "t + ?n \ l + ?n" by auto + ultimately have "((take t sts @ sts0) @ drop t sts)[l + ?n := v] = + (take t sts @ sts0) @ (drop t sts)[(l + ?n - (t + ?n)) := v]" + by (smt list_update_append replicate_nth) + with eq_take_drop_t + show ?thesis by auto +qed + +lemma l_n_v_s1: + assumes "l < length sts" + and "\ t \ l" + shows "(take t sts @ sts0 @ drop t sts)[l := v] = + take t (sts[l := v]) @ sts0 @ drop t (sts[l := v])" +proof(cases "t < length sts") + case False + hence h: "take t sts = sts" "drop t sts = []" + "take t (sts[l:=v]) = sts [l:=v]" + "drop t (sts[l:=v]) = []" + by auto + with assms(1) + show ?thesis + by (metis list_update_append) +next + case True + with assms(2) + have h: "(take t sts)[l:=v] = take t (sts[l:=v])" + "(sts[l:=v]) = (take t sts)[l:=v]@drop t sts" + "length (take t sts) = t" + apply (smt length_list_update length_take nth_equalityI nth_list_update nth_take) + apply (smt True append_take_drop_id assms(2) length_take list_update_append1) + by (smt True length_take) + from h(2,3) have "drop t (sts[l := v]) = drop t sts" + by (metis append_eq_conv_conj length_list_update) + with h(1) + show ?thesis + apply simp + by (metis assms(2) h(3) list_update_append1 not_leE) +qed + +lemma l_n_v_s2: + assumes "l0 < length env" + and "t \ l0" + and "\ t \ l1" + shows "(take t env @ es @ drop t env) ! l1 = env ! l1" +proof - + from assms(1, 2) have "t < length env" by auto + hence h: "env = take t env @ drop t env" + "length (take t env) = t" + apply (metis append_take_drop_id) + by (smt `t < length env` length_take) + with assms(3) show ?thesis + by (smt nth_append) +qed + +lemma l_n_v_s3: + assumes "l0 < length env" + and "\ t \ l0" + shows "(take t env @ es @ drop t env) ! l0 = env ! l0" +proof(cases "t < length env") + case True + hence h: "env = take t env @ drop t env" + "length (take t env) = t" + apply (metis append_take_drop_id) + by (smt `t < length env` length_take) + with assms(2) show ?thesis + by (smt nth_append) +next + case False + hence "take t env = env" by auto + with assms(1) show ?thesis + by (metis nth_append) +qed + +subsection {* Invariant under lifts and perms *} + +definition "lift_b t i j = (if (j \ t) then (j + i) else j)" + +fun lift_t :: "nat \ nat \ cpg \ cpg" +where "lift_t t i (CInstr ((act0, l0), (act1, l1))) = + (CInstr ((act0, lift_b t i (nat_of l0)), + (act1, lift_b t i (nat_of l1))))" | + "lift_t t i (CLabel l) = CLabel (lift_b t i l)" | + "lift_t t i (CSeq c1 c2) = CSeq (lift_t t i c1) (lift_t t i c2)" | + "lift_t t i (CLocal c) = CLocal (lift_t (t + 1) i c)" + +definition "lift0 (i::nat) cpg = lift_t 0 i cpg" + +definition "perm_b t i j k = (if ((k::nat) = i \ i < t \ j < t) then j else + if (k = j \ i < t \ j < t) then i else k)" + +fun perm :: "nat \ nat \ nat \ cpg \ cpg" +where "perm t i j (CInstr ((act0, l0), (act1, l1))) = + (CInstr ((act0, perm_b t i j l0), (act1, perm_b t i j l1)))" | + "perm t i j (CLabel l) = CLabel (perm_b t i j l)" | + "perm t i j (CSeq c1 c2) = CSeq (perm t i j c1) (perm t i j c2)" | + "perm t i j (CLocal c) = CLocal (perm (t + 1) (i + 1) (j + 1) c)" + +definition "map_idx f sts = map (\ k. sts!(f (nat k))) [0 .. int (length sts) - 1]" + +definition "perm_s i j sts = map_idx (perm_b (length sts) i j) sts" + +fun lift_es :: "(tstate list \ nat) list \ tstate list \ tstate list" where + "lift_es [] env = env" + | "lift_es ((env', t)#ops) env = lift_es ops (take t env @ env' @ drop t env)" + +fun lift_ss :: "(tstate list \ nat) list \ status list \ status list" where + "lift_ss [] sts = sts" + | "lift_ss ((env', t)#ops) sts = lift_ss ops (take t sts @ map (\ x. Free) env' @ drop t sts)" + + +fun lift_ts :: "(nat \ nat) list \ cpg \ cpg" where + "lift_ts [] cpg = cpg" + | "lift_ts ((lenv, t)#ops) cpg = lift_ts ops (lift_t t lenv cpg)" + +fun perm_ss :: "(nat \ nat) list \ 'a list \ 'a list" where + "perm_ss [] env = env" + | "perm_ss ((i, j)#ops) env = perm_ss ops (perm_s i j env)" + +fun perms :: "nat => (nat \ nat) list \ cpg \ cpg" where + "perms n [] cpg = cpg" + | "perms n ((i, j)#ops) cpg = perms n ops (perm n i j cpg)" + +definition + "adjust_cpg len sps lfs cpg = lift_ts lfs (perms len sps cpg)" + +definition + "red_lfs lfs = map (apfst length) lfs" + +definition + "adjust_env sps lfs env = lift_es lfs (perm_ss sps env)" + +definition + "adjust_sts sps lfs sts = lift_ss lfs (perm_ss sps sts)" + +fun sts_disj_test :: "status list \ status list \ bool" where + "sts_disj_test [] [] = True" + | "sts_disj_test [] (s#ss) = False" + | "sts_disj_test (s#ss) [] = False" + | "sts_disj_test (s1#ss1) (s2#ss2) = (case (s1, s2) of + (Bound, Bound) \ False + | _ \ sts_disj_test ss1 ss2)" + +lemma inj_perm_b: "inj (perm_b t i j)" +proof(induct rule:injI) + case (1 x y) + thus ?case + by (unfold perm_b_def, auto split:if_splits) +qed + +lemma lift_wf_cpg_test: + assumes "wf_cpg_test sts cpg = (True, sts')" + shows "wf_cpg_test (take t sts @ sts0 @ drop t sts) (lift_t t (length sts0) cpg) = + (True, take t sts' @ sts0 @ drop t sts')" + using assms +proof(induct cpg arbitrary:t sts0 sts sts') + case (CInstr instr t sts0 sts sts') + obtain a0 l0 a1 l1 where eq_instr: "instr = ((a0, St l0), (a1, St l1))" + by (metis surj_pair tstate.exhaust) + from CInstr + show ?case + by (auto simp:eq_instr lift_b_def) +next + case (CLabel l t sts0 sts sts') + thus ?case + apply (auto simp:lift_b_def + replicate_nth replicate_update l_n_v_orig l_n_v_s) + apply (metis (mono_tags) diff_diff_cancel length_drop length_rev + linear not_less nth_append nth_take rev_take take_all) + by (simp add:l_n_v_s1) +next + case (CSeq c1 c2 sts0 sts sts') + thus ?case + by (auto simp: lift0_def lift_b_def split:prod.splits) +next + case (CLocal body t sts0 sts sts') + from this(2) obtain sts1 where h: "wf_cpg_test (Free # sts) body = (True, sts1)" "tl sts1 = sts'" + by (auto simp:lift0_def lift_b_def split:prod.splits) + let ?lift_s = "\ t sts. take t sts @ sts0 @ drop t sts" + have eq_lift_1: "(?lift_s (Suc t) (Free # sts)) = Free#?lift_s t sts" + by (simp) + from wf_cpg_test_le[OF h(1)] have "length (Free#sts) = length sts1" + by (unfold less_eq_list_def, simp) + hence eq_sts1: "sts1 = hd sts1 # tl sts1" + by (metis append_Nil append_eq_conv_conj hd.simps list.exhaust tl.simps(2)) + from CLocal(1)[OF h(1), of "Suc t", of "sts0", unfolded eq_lift_1] + show ?case + apply (simp, subst eq_sts1, simp) + apply (simp add:h(2)) + by (subst eq_sts1, simp add:h(2)) +qed + +lemma lift_c2t: + assumes "wf_cpg_test sts cpg = (True, sts')" + and "length env = length sts" + shows "c2t (take t env @ es @ drop t env) (lift_t t (length es) cpg) = + c2t env cpg" + using assms +proof(induct cpg arbitrary: t env es sts sts') + case (CInstr instr t env es sts sts') + obtain a0 l0 a1 l1 where eq_instr: "instr = ((a0, St l0), (a1, St l1))" + by (metis nat_of.cases surj_pair) + from CInstr have h: "l0 < length env" "l1 < length env" + by (auto simp:eq_instr) + with CInstr(2) + show ?case + by (auto simp:eq_instr lift_b_def l_n_v_orig l_n_v_s2 l_n_v_s3) +next + case (CLabel l t env es sts sts') + thus ?case + by (auto simp:lift_b_def + replicate_nth replicate_update l_n_v_orig l_n_v_s3) +next + case (CSeq c1 c2 t env es sts sts') + from CSeq(3) obtain sts1 + where h: "wf_cpg_test sts c1 = (True, sts1)" "wf_cpg_test sts1 c2 = (True, sts')" + by (auto split:prod.splits) + from wf_cpg_test_le[OF h(1)] have "length sts = length sts1" + by (auto simp:less_eq_list_def) + from CSeq(4)[unfolded this] have eq_len_env: "length env = length sts1" . + from CSeq(1)[OF h(1) CSeq(4)] + CSeq(2)[OF h(2) eq_len_env] + show ?case + by (auto simp: lift0_def lift_b_def split:prod.splits) +next + case (CLocal body t env es sts sts') + { fix x + from CLocal(2) + obtain sts1 where h1: "wf_cpg_test (Free # sts) body = (True, sts1)" + by (auto split:prod.splits) + from CLocal(3) have "length (x#env) = length (Free # sts)" by simp + from CLocal(1)[OF h1 this, of "Suc t"] + have "c2t (x # take t env @ es @ drop t env) (lift_t (Suc t) (length es) body) = + c2t (x # env) body" + by simp + } thus ?case by simp +qed + +lemma upto_len: "length [i .. j] = (if j < i then 0 else (nat (j - i + 1)))" +proof(induct i j rule:upto.induct) + case (1 i j) + show ?case + proof(cases "j < i") + case True + thus ?thesis by simp + next + case False + hence eq_ij: "[i..j] = i # [i + 1..j]" by (simp add:upto.simps) + from 1 False + show ?thesis + by (auto simp:eq_ij) + qed +qed + +lemma upto_append: + assumes "x \ y + 1" + shows "[x .. y + 1] = [x .. y]@[y + 1]" + using assms + by (induct x y rule:upto.induct, auto simp:upto.simps) + +lemma nth_upto: + assumes "l < length sts" + shows "[0..(int (length sts)) - 1]!l = int l" + using assms +proof(induct sts arbitrary:l) + case Nil + thus ?case by simp +next + case (Cons s sts l) + from Cons(2) + have "0 \ (int (length sts) - 1) + 1" by auto + from upto_append[OF this] + have eq_upto: "[0..int (length sts)] = [0..int (length sts) - 1] @ [int (length sts)]" + by simp + show ?case + proof(cases "l < length sts") + case True + with Cons(1)[OF True] eq_upto + show ?thesis + apply simp + by (smt nth_append take_eq_Nil upto_len) + next + case False + with Cons(2) have eq_l: "l = length sts" by simp + show ?thesis + proof(cases sts) + case (Cons x xs) + have "[0..1 + int (length xs)] = [0 .. int (length xs)]@[1 + int (length xs)]" + by (smt upto_append) + moreover have "length [0 .. int (length xs)] = Suc (length xs)" + by (smt upto_len) + moreover note Cons + ultimately show ?thesis + apply (simp add:eq_l) + by (smt nth_Cons' nth_append) + qed (simp add:upto_len upto.simps eq_l) + qed +qed + +lemma map_idx_idx: + assumes "l < length sts" + shows "(map_idx f sts)!l = sts!(f l)" +proof - + from assms have lt_l: "l < length [0..int (length sts) - 1]" + by (smt upto_len) + show ?thesis + apply (unfold map_idx_def nth_map[OF lt_l]) + by (metis assms nat_int nth_upto) +qed + +lemma map_idx_len: "length (map_idx f sts) = length sts" + apply (unfold map_idx_def) + by (smt length_map upto_len) + +lemma map_idx_eq: + assumes "\ l < length sts. f l = g l" + shows "map_idx f sts = map_idx g sts" +proof(induct rule: nth_equalityI) + case 1 + show "length (map_idx f sts) = length (map_idx g sts)" + by (metis map_idx_len) +next + case 2 + { fix l + assume "l < length (map_idx f sts)" + hence "l < length sts" + by (metis map_idx_len) + from map_idx_idx[OF this] and assms and this + have "map_idx f sts ! l = map_idx g sts ! l" + by (smt list_eq_iff_nth_eq map_idx_idx map_idx_len) + } thus ?case by auto +qed + +lemma perm_s_commut: "perm_s i j sts = perm_s j i sts" + apply (unfold perm_s_def, rule map_idx_eq, unfold perm_b_def) + by smt + +lemma map_idx_id: "map_idx id sts = sts" +proof(induct rule:nth_equalityI) + case 1 + from map_idx_len show "length (map_idx id sts) = length sts" . +next + case 2 + { fix l + assume "l < length (map_idx id sts)" + from map_idx_idx[OF this[unfolded map_idx_len]] + have "map_idx id sts ! l = sts ! l" by simp + } thus ?case by auto +qed + +lemma perm_s_lt_i: + assumes "\ i < length sts" + shows "perm_s i j sts = sts" +proof - + have "map_idx (perm_b (length sts) i j) sts = map_idx id sts" + proof(rule map_idx_eq, default, clarsimp) + fix l + assume "l < length sts" + with assms + show "perm_b (length sts) i j l = l" + by (unfold perm_b_def, auto) + qed + with map_idx_id + have "map_idx (perm_b (length sts) i j) sts = sts" by simp + thus ?thesis by (simp add:perm_s_def) +qed + +lemma perm_s_lt: + assumes "\ i < length sts \ \ j < length sts" + shows "perm_s i j sts = sts" + using assms +proof + assume "\ i < length sts" + from perm_s_lt_i[OF this] show ?thesis . +next + assume "\ j < length sts" + from perm_s_lt_i[OF this, of i, unfolded perm_s_commut] + show ?thesis . +qed + +lemma perm_s_update_i: + assumes "i < length sts" + and "j < length sts" + shows "sts ! i = perm_s i j sts ! j" +proof - + from map_idx_idx[OF assms(2)] + have "map_idx (perm_b (length sts) i j) sts ! j = sts!(perm_b (length sts) i j j)" . + with assms + show ?thesis + by (auto simp:perm_s_def perm_b_def) +qed + +lemma nth_perm_s_neq: + assumes "l \ j" + and "l \ i" + and "l < length sts" + shows "sts ! l = perm_s i j sts ! l" +proof - + have "map_idx (perm_b (length sts) i j) sts ! l = sts!(perm_b (length sts) i j l)" + by (metis assms(3) map_idx_def map_idx_idx) + with assms + show ?thesis + by (unfold perm_s_def perm_b_def, auto) +qed + +lemma map_idx_update: + assumes "f j = i" + and "inj f" + and "i < length sts" + and "j < length sts" + shows "map_idx f (sts[i:=v]) = map_idx f sts[j := v]" +proof(induct rule:nth_equalityI) + case 1 + show "length (map_idx f (sts[i := v])) = length (map_idx f sts[j := v])" + by (metis length_list_update map_idx_len) +next + case 2 + { fix l + assume lt_l: "l < length (map_idx f (sts[i := v]))" + have eq_nth: "sts[i := v] ! f l = map_idx f sts[j := v] ! l" + proof(cases "f l = i") + case False + from lt_l have "l < length sts" + by (metis length_list_update map_idx_len) + from map_idx_idx[OF this, of f] have " map_idx f sts ! l = sts ! f l" . + moreover from False assms have "l \ j" by auto + moreover note False + ultimately show ?thesis by simp + next + case True + with assms have eq_l: "l = j" + by (metis inj_eq) + moreover from lt_l eq_l + have "j < length (map_idx f sts[j := v])" + by (metis length_list_update map_idx_len) + moreover note True assms + ultimately show ?thesis by simp + qed + from lt_l have "l < length (sts[i := v])" + by (metis map_idx_len) + from map_idx_idx[OF this] eq_nth + have "map_idx f (sts[i := v]) ! l = map_idx f sts[j := v] ! l" by simp + } thus ?case by auto +qed + +lemma perm_s_update: + assumes "i < length sts" + and "j < length sts" + shows "(perm_s i j sts)[i := v] = perm_s i j (sts[j := v])" +proof - + have "map_idx (perm_b (length (sts[j := v])) i j) (sts[j := v]) = + map_idx (perm_b (length (sts[j := v])) i j) sts[i := v]" + proof(rule map_idx_update[OF _ _ assms(2, 1)]) + from inj_perm_b show "inj (perm_b (length (sts[j := v])) i j)" . + next + from assms show "perm_b (length (sts[j := v])) i j i = j" + by (auto simp:perm_b_def) + qed + hence "map_idx (perm_b (length sts) i j) sts[i := v] = + map_idx (perm_b (length (sts[j := v])) i j) (sts[j := v])" + by simp + thus ?thesis by (simp add:perm_s_def) +qed + +lemma perm_s_len: "length (perm_s i j sts') = length sts'" + apply (unfold perm_s_def map_idx_def) + by (smt Nil_is_map_conv length_0_conv length_greater_0_conv length_map neq_if_length_neq upto_len) + +lemma perm_s_update_neq: + assumes "l \ i" + and "l \ j" + shows "perm_s i j sts[l := v] = perm_s i j (sts[l := v])" +proof(cases "i < length sts \ j < length sts") + case False + with perm_s_lt have "perm_s i j sts = sts" by auto + moreover have "perm_s i j (sts[l:=v]) = sts[l:=v]" + proof - + have "length (sts[l:=v]) = length sts" by auto + from False[folded this] perm_s_lt + show ?thesis by metis + qed + ultimately show ?thesis by simp +next + case True + note lt_ij = this + show ?thesis + proof(cases "l < length sts") + case False + hence "sts[l:=v] = sts" by auto + moreover have "perm_s i j sts[l := v] = perm_s i j sts" + proof - + from False and perm_s_len + have "\ l < length (perm_s i j sts)" by metis + thus ?thesis by auto + qed + ultimately show ?thesis by simp + next + case True + show ?thesis + proof - + have "map_idx (perm_b (length (sts[l := v])) i j) (sts[l := v]) = + map_idx (perm_b (length (sts[l := v])) i j) sts[l := v]" + proof(induct rule:map_idx_update [OF _ inj_perm_b True True]) + case 1 + from assms lt_ij + show ?case + by (unfold perm_b_def, auto) + qed + thus ?thesis + by (unfold perm_s_def, simp) + qed + qed +qed + +lemma perm_sb: "(perm_s i j sts)[perm_b (length sts) i j l := v] = perm_s i j (sts[l := v])" + apply(subst perm_b_def, auto simp:perm_s_len perm_s_lt perm_s_update) + apply (subst perm_s_commut, subst (2) perm_s_commut, rule_tac perm_s_update, auto) + by (rule_tac perm_s_update_neq, auto) + +lemma perm_s_id: "perm_s i i sts = sts" (is "?L = ?R") +proof - + from map_idx_id have "?R = map_idx id sts" by metis + also have "\ = ?L" + by (unfold perm_s_def, rule map_idx_eq, unfold perm_b_def, auto) + finally show ?thesis by simp +qed + +lemma upto_map: + assumes "i \ j" + shows "[i .. j] = i # map (\ x. x + 1) [i .. (j - 1)]" + using assms +proof(induct i j rule:upto.induct) + case (1 i j) + show ?case (is "?L = ?R") + proof - + from 1(2) + have eq_l: "?L = i # [i+1 .. j]" by (simp add:upto.simps) + show ?thesis + proof(cases "i + 1 \ j") + case False + with eq_l show ?thesis by (auto simp:upto.simps) + next + case True + have "[i + 1..j] = map (\x. x + 1) [i..j - 1]" + by (smt "1.hyps" Cons_eq_map_conv True upto.simps) + with eq_l + show ?thesis by simp + qed + qed +qed + +lemma perm_s_cons: "(perm_s (Suc i) (Suc j) (s # sts)) = s#perm_s i j sts" +proof - + have le_0: "0 \ int (length (s # sts)) - 1" by simp + have "map (\k. (s # sts) ! perm_b (length (s # sts)) (Suc i) (Suc j) (nat k)) + [0..int (length (s # sts)) - 1] = + s # map (\k. sts ! perm_b (length sts) i j (nat k)) [0..int (length sts) - 1]" + by (unfold upto_map[OF le_0], auto simp:perm_b_def, smt+) + thus ?thesis by (unfold perm_s_def map_idx_def, simp) +qed + +lemma perm_wf_cpg_test: + assumes "wf_cpg_test sts cpg = (True, sts')" + shows "wf_cpg_test (perm_s i j sts) (perm (length sts) i j cpg) = + (True, perm_s i j sts')" + using assms +proof(induct cpg arbitrary:t i j sts sts') + case (CInstr instr i j sts sts') + obtain a0 l0 a1 l1 where eq_instr: "instr = ((a0, St l0), (a1, St l1))" + by (metis surj_pair tstate.exhaust) + from CInstr + show ?case + apply (unfold eq_instr, clarsimp) + by (unfold perm_s_len perm_b_def, clarsimp) +next + case (CLabel l i j sts sts') + have "(perm_s i j sts)[perm_b (length sts) i j l := Bound] = perm_s i j (sts[l := Bound])" + by (metis perm_sb) + with CLabel + show ?case + apply (auto simp:perm_s_len perm_sb) + apply (subst perm_b_def, auto simp:perm_sb) + apply (subst perm_b_def, auto simp:perm_s_lt perm_s_update_i) + apply (unfold perm_s_id, subst perm_s_commut, simp add: perm_s_update_i[symmetric]) + apply (simp add:perm_s_update_i[symmetric]) + by (simp add: nth_perm_s_neq[symmetric]) +next + case (CSeq c1 c2 i j sts sts') + thus ?case + apply (auto split:prod.splits) + apply (metis (hide_lams, no_types) less_eq_list_def prod.inject wf_cpg_test_le) + by (metis (hide_lams, no_types) less_eq_list_def prod.inject wf_cpg_test_le) +next + case (CLocal body i j sts sts') + from this(2) obtain sts1 where h: "wf_cpg_test (Free # sts) body = (True, sts1)" "tl sts1 = sts'" + by (auto simp:lift0_def lift_b_def split:prod.splits) + from wf_cpg_test_le[OF h(1)] have "length (Free#sts) = length sts1" + by (unfold less_eq_list_def, simp) + hence eq_sts1: "sts1 = hd sts1 # tl sts1" + by (metis append_Nil append_eq_conv_conj hd.simps list.exhaust tl.simps(2)) + from CLocal(1)[OF h(1), of "Suc i" "Suc j"] h(2) eq_sts1 + show ?case + apply (auto split:prod.splits simp:perm_s_cons) + by (metis perm_s_cons tl.simps(2)) +qed + +lemma nth_perm_sb: + assumes "l0 < length env" + shows "perm_s i j env ! perm_b (length env) i j l0 = env ! l0" + by (metis assms nth_perm_s_neq perm_b_def perm_s_commut perm_s_lt perm_s_update_i) + + +lemma perm_c2t: + assumes "wf_cpg_test sts cpg = (True, sts')" + and "length env = length sts" + shows "c2t (perm_s i j env) (perm (length env) i j cpg) = + c2t env cpg" + using assms +proof(induct cpg arbitrary:i j env sts sts') + case (CInstr instr i j env sts sts') + obtain a0 l0 a1 l1 where eq_instr: "instr = ((a0, St l0), (a1, St l1))" + by (metis surj_pair tstate.exhaust) + from CInstr have h: "l0 < length env" "l1 < length env" + by (auto simp:eq_instr) + with CInstr(2) + show ?case + apply (auto simp:eq_instr) + by (metis nth_perm_sb)+ +next + case (CLabel l t env es sts sts') + thus ?case + apply (auto) + by (metis nth_perm_sb) +next + case (CSeq c1 c2 i j env sts sts') + from CSeq(3) obtain sts1 + where h: "wf_cpg_test sts c1 = (True, sts1)" "wf_cpg_test sts1 c2 = (True, sts')" + by (auto split:prod.splits) + from wf_cpg_test_le[OF h(1)] have "length sts = length sts1" + by (auto simp:less_eq_list_def) + from CSeq(4)[unfolded this] have eq_len_env: "length env = length sts1" . + from CSeq(1)[OF h(1) CSeq(4)] + CSeq(2)[OF h(2) eq_len_env] + show ?case by auto +next + case (CLocal body i j env sts sts') + { fix x + from CLocal(2, 3) + obtain sts1 where "wf_cpg_test (Free # sts) body = (True, sts1)" + "length (x#env) = length (Free # sts)" + by (auto split:prod.splits) + from CLocal(1)[OF this] + have "(c2t (x # perm_s i j env) (perm (Suc (length env)) (Suc i) (Suc j) body)) = + (c2t (x # env) body)" + by (metis Suc_length_conv perm_s_cons) + } thus ?case by simp +qed + +lemma wf_cpg_test_disj_aux1: + assumes "sts_disj sts1 (sts[l := Bound] - sts)" + "l < length sts" + "sts ! l = Free" + shows "(sts1 + sts) ! l = Free" +proof - + from assms(1)[unfolded sts_disj_def] + have h: "length sts1 = length (sts[l := Bound] - sts)" + "(\i (sts1 ! i = Bound \ (sts[l := Bound] - sts) ! i = Bound))" + by auto + from h(1) assms(2) + have lt_l: "l < length sts1" + "l < length (sts[l := Bound] - sts)" + "l < length (sts1 + sts)" + apply (smt length_list_update minus_list_len) + apply (smt assms(2) length_list_update minus_list_len) + by (smt assms(2) h(1) length_list_update length_sts_plus minus_list_len) + from h(2)[rule_format, of l, OF this(1)] + have " \ (sts1 ! l = Bound \ (sts[l := Bound] - sts) ! l = Bound)" . + with assms(3) nth_sts_minus[OF lt_l(2)] nth_sts_plus[OF lt_l(3)] assms(2) + show ?thesis + by (cases "sts1!l", auto) +qed + +lemma wf_cpg_test_disj_aux2: + assumes "sts_disj sts1 (sts[l := Bound] - sts)" + " l < length sts" + shows "(sts1 + sts)[l := Bound] = sts1 + sts[l := Bound]" +proof - + from assms have lt_l: "l < length (sts1 + sts[l:=Bound])" + "l < length (sts1 + sts)" + apply (smt length_list_update length_sts_plus minus_list_len sts_disj_def) + by (smt assms(1) assms(2) length_list_update length_sts_plus minus_list_len sts_disj_def) + show ?thesis + proof(induct rule:nth_equalityI) + case 1 + show ?case + by (smt assms(1) length_list_update length_sts_plus minus_list_len sts_disj_def) + next + case 2 + { fix i + assume lt_i: "i < length ((sts1 + sts)[l := Bound])" + have " (sts1 + sts)[l := Bound] ! i = (sts1 + sts[l := Bound]) ! i" + proof(cases "i = l") + case True + with nth_sts_plus[OF lt_l(1)] assms(2) nth_sts_plus[OF lt_l(2)] lt_l + show ?thesis + by (cases "sts1 ! l", auto) + next + case False + from lt_i have "i < length (sts1 + sts)" "i < length (sts1 + sts[l := Bound])" + apply auto + by (metis length_list_update plus_list_len) + from nth_sts_plus[OF this(1)] nth_sts_plus[OF this(2)] lt_i lt_l False + show ?thesis + by simp + qed + } thus ?case by auto + qed +qed + +lemma sts_list_plus_commut: + shows "sts1 + sts2 = sts2 + (sts1:: status list)" +proof(induct rule:nth_equalityI) + case 1 + show ?case + by (metis min_max.inf.commute plus_list_len) +next + case 2 + { fix i + assume lt_i1: "ii (sts1 ! i = Bound \ sts2 ! i = Bound))" + by (unfold sts_disj_def, auto) + from h(1) have "length (Free # sts1) = length (s # sts2)" by simp + moreover { + fix i + assume lt_i: "i ((Free # sts1) ! i = Bound \ (s # sts2) ! i = Bound)" + proof(cases "i") + case 0 + thus ?thesis by simp + next + case (Suc k) + from h(2)[rule_format, of k] lt_i[unfolded Suc] Suc + show ?thesis by auto + qed + } + ultimately show ?thesis by (auto simp:sts_disj_def) +qed + +lemma sts_disj_uncomb: + assumes "sts_disj sts (sts1 + sts2)" + and "sts_disj sts1 sts2" + shows "sts_disj sts sts1" "sts_disj sts sts2" + using assms + apply (smt assms(1) assms(2) length_sts_plus nth_sts_plus plus_status.simps(2) sts_disj_def) + by (smt assms(1) assms(2) length_sts_plus nth_sts_plus + plus_status.simps(2) sts_disj_def sts_list_plus_commut) + +lemma wf_cpg_test_disj: + assumes "wf_cpg_test sts cpg = (True, sts')" + and "sts_disj sts1 (sts' - sts)" + shows "wf_cpg_test (sts1 + sts) cpg = (True, sts1 + sts')" + using assms +proof(induct cpg arbitrary:sts sts1 sts') + case (CInstr instr sts sts1 sts') + obtain a0 l0 a1 l1 where eq_instr: "instr = ((a0, St l0), (a1, St l1))" + by (metis nat_of.cases surj_pair) + with CInstr(1) have h: "l0 < length sts" "l1 < length sts" "sts = sts'" by auto + with CInstr eq_instr + show ?case + apply (auto) + by (smt length_sts_plus minus_list_len sts_disj_def)+ +next + case (CLabel l sts sts1 sts') + thus ?case + apply auto + apply (smt length_list_update length_sts_plus minus_list_len sts_disj_def) + by (auto simp: wf_cpg_test_disj_aux1 wf_cpg_test_disj_aux2) +next + case (CSeq c1 c2 sts sts1 sts') + from CSeq(3) obtain sts'' + where h: "wf_cpg_test sts c1 = (True, sts'')" "wf_cpg_test sts'' c2 = (True, sts')" + by (auto split:prod.splits) + from wf_cpg_test_le[OF h(1)] have "length sts = length sts''" + by (auto simp:less_eq_list_def) + from sts_le_comb[OF wf_cpg_test_le[OF h(1)] wf_cpg_test_le[OF h(2)]] + have " sts' - sts = (sts'' - sts) + (sts' - sts'')" "sts_disj (sts'' - sts) (sts' - sts'')" + by auto + from sts_disj_uncomb[OF CSeq(4)[unfolded this(1)] this(2)] + have "sts_disj sts1 (sts'' - sts)" "sts_disj sts1 (sts' - sts'')" . + from CSeq(1)[OF h(1) this(1)] CSeq(2)[OF h(2) this(2)] + have "wf_cpg_test (sts1 + sts) c1 = (True, sts1 + sts'')" + "wf_cpg_test (sts1 + sts'') c2 = (True, sts1 + sts')" . + thus ?case + by simp +next + case (CLocal body sts sts1 sts') + from this(2) + obtain sts'' where h: "wf_cpg_test (Free # sts) body = (True, sts'')" "sts' = tl sts''" + by (auto split:prod.splits) + from wf_cpg_test_le[OF h(1), unfolded less_eq_list_def] h(2) + obtain s where eq_sts'': "sts'' = s#sts'" + by (metis Suc_length_conv list.size(4) tl.simps(2)) + let ?sts = "Free#sts1" + from CLocal(3) have "sts_disj ?sts (sts'' - (Free # sts))" + apply (unfold eq_sts'', simp) + by (metis sts_disj_cons) + from CLocal(1)[OF h(1) this] eq_sts'' + show ?case + by (auto split:prod.splits) +qed + +lemma sts_disj_free: + assumes "list_all (op = Free) sts" + and "length sts' = length sts" + shows "sts_disj sts' sts" +by (metis (full_types) assms(1) assms(2) list_all_length + status.distinct(1) sts_disj_def) + +lemma all_free_plus: + assumes "length sts' = length sts" + and "list_all (op = Free) sts" + shows "sts' + sts = sts'" + using assms +proof(induct sts' arbitrary:sts) + case (Cons s sts' sts) + note cs = Cons + thus ?case + proof(cases "sts") + case (Cons s1 sts1) + with cs + show ?thesis + by (smt list.size(4) list_all_simps(1) + plus_list.simps(3) plus_status.simps(1) sts_list_plus_commut) + qed auto +qed auto + +lemma wf_cpg_test_extrapo: + assumes "wf_cpg_test sts cpg = (True, sts)" + and "list_all (op = Free) sts" + and "length sts' = length sts" + shows "wf_cpg_test sts' cpg = (True, sts')" +proof - + have "sts_disj sts' (sts - sts)" + proof(rule sts_disj_free) + from assms(2) + show "list_all (op = Free) (sts - sts)" + by (induct sts, auto) + next + from assms(3) show "length sts' = length (sts - sts)" + by (metis length_sts_plus minus_list_len plus_list_len) + qed + from wf_cpg_test_disj [OF assms(1) this] + have "wf_cpg_test (sts' + sts) cpg = (True, sts' + sts)" . + moreover from all_free_plus[OF assms(3, 2)] have "sts' + sts = sts'" . + finally show ?thesis by simp +qed + +lemma perms_wf_cpg_test: + assumes "wf_cpg_test sts cpg = (True, sts')" + shows "wf_cpg_test (perm_ss ops sts) (perms (length sts) ops cpg) = + (True, perm_ss ops sts')" + using assms +proof(induct ops arbitrary:sts cpg sts') + case (Cons sp ops sts cpg sts') + show ?case + proof(cases "sp") + case (Pair i j) + show ?thesis + proof - + let ?sts = "(perm_s i j sts)" and ?cpg = "(perm (length sts) i j cpg)" + and ?sts' = "perm_s i j sts'" + have "wf_cpg_test (perm_ss ops ?sts) (perms (length ?sts) ops ?cpg) = + (True, perm_ss ops ?sts')" + proof(rule Cons(1)) + show "wf_cpg_test (perm_s i j sts) (perm (length sts) i j cpg) = (True, perm_s i j sts')" + by (metis Cons.prems perm_wf_cpg_test) + qed + thus ?thesis + apply (unfold Pair) + apply simp + by (metis perm_s_len) + qed + qed +qed auto + +lemma perm_ss_len: "length (perm_ss ops xs) = length xs" +proof(induct ops arbitrary:xs) + case (Cons sp ops xs) + show ?case + proof(cases "sp") + case (Pair i j) + show ?thesis + proof - + let ?xs = "(perm_s i j xs)" + have "length (perm_ss ops ?xs) = length ?xs" + by (metis Cons.hyps) + also have "\ = length xs" + by (metis perm_s_len) + finally show ?thesis + by (unfold Pair, simp) + qed + qed +qed auto + +lemma perms_c2t: + assumes "wf_cpg_test sts cpg = (True, sts')" + and "length env = length sts" + shows "c2t (perm_ss ops env) (perms (length env) ops cpg) = c2t env cpg" + using assms +proof(induct ops arbitrary:sts cpg sts' env) + case (Cons sp ops sts cpg sts' env) + show ?case + proof(cases "sp") + case (Pair i j) + show ?thesis + proof - + let ?env = "(perm_s i j env)" and ?cpg = "(perm (length env) i j cpg)" + have " c2t (perm_ss ops ?env) (perms (length ?env) ops ?cpg) = c2t ?env ?cpg" + proof(rule Cons(1)) + from perm_wf_cpg_test[OF Cons(2), of i j, folded Cons(3)] + show "wf_cpg_test (perm_s i j sts) (perm (length env) i j cpg) = (True, perm_s i j sts')" . + next + show "length (perm_s i j env) = length (perm_s i j sts)" + by (metis Cons.prems(2) perm_s_len) + qed + also have "\ = c2t env cpg" + by (metis Cons.prems(1) Cons.prems(2) perm_c2t) + finally show ?thesis + apply (unfold Pair) + apply simp + by (metis perm_s_len) + qed + qed +qed auto + +lemma red_lfs_nil: "red_lfs [] = []" + by (simp add:red_lfs_def) + +lemma red_lfs_cons: "red_lfs ((env, t)#lfs) = (length env, t)#(red_lfs lfs)" + by (simp add:red_lfs_def) + +lemmas red_lfs_simps [simp] = red_lfs_nil red_lfs_cons + +lemma lifts_wf_cpg_test: + assumes "wf_cpg_test sts cpg = (True, sts')" + shows "wf_cpg_test (lift_ss ops sts) (lift_ts (red_lfs ops) cpg) + = (True, lift_ss ops sts')" + using assms +proof(induct ops arbitrary:sts cpg sts') + case (Cons sp ops sts cpg sts') + show ?case + proof(cases "sp") + case (Pair env' t) + thus ?thesis + proof - + let ?sts = "(take t sts @ map (\x. Free) env' @ drop t sts)" + and ?sts' = "(take t sts' @ map (\x. Free) env' @ drop t sts')" + and ?cpg = "(lift_t t (length env') cpg)" + have "wf_cpg_test (lift_ss ops ?sts) (lift_ts (red_lfs ops) ?cpg) = (True, lift_ss ops ?sts')" + proof(induct rule: Cons(1)) + case 1 + show ?case + by (metis Cons.prems length_map lift_wf_cpg_test) + qed + thus ?thesis + by (unfold Pair, simp) + qed + qed +qed auto + +lemma lifts_c2t: + assumes "wf_cpg_test sts cpg = (True, sts')" + and "length env = length sts" + shows "c2t (lift_es ops env) (lift_ts (red_lfs ops) cpg) = c2t env cpg" + using assms +proof(induct ops arbitrary:sts cpg sts' env) + case (Cons sp ops sts cpg sts' env) + show ?case + proof(cases "sp") + case (Pair env' t) + show ?thesis + proof - + let ?env = "(take t env @ env' @ drop t env)" + and ?cpg = "(lift_t t (length env') cpg)" + have "c2t (lift_es ops ?env) (lift_ts (red_lfs ops) ?cpg) = c2t ?env ?cpg" + proof(rule Cons(1)) + from lift_wf_cpg_test[OF Cons(2), of t "map (\ x. Free) env'", simplified length_map] + show "wf_cpg_test (take t sts @ map (\x. Free) env' @ drop t sts) + (lift_t t (length env') cpg) = + (True, take t sts' @ map (\x. Free) env' @ drop t sts')" . + next + show "length (take t env @ env' @ drop t env) = + length (take t sts @ map (\x. Free) env' @ drop t sts)" + by (metis (full_types) Cons.prems(2) Pair assms(2) length_append + length_drop length_map length_take) + qed + also have "\ = c2t env cpg" + by (metis Cons.prems(1) Cons.prems(2) lift_c2t) + finally show ?thesis + by (unfold Pair, simp) + qed + qed +qed auto + +lemma adjust_c2t: + assumes "wf_cpg_test sts cpg = (True, sts')" + and "length env = length sts" + shows "c2t (adjust_env sps lfs env) (adjust_cpg (length sts) sps (red_lfs lfs) cpg) = c2t env cpg" +proof - + let ?cpg = "(perms (length sts) sps cpg)" + and ?env = "(perm_ss sps env)" + have "c2t (lift_es lfs ?env) + (lift_ts (red_lfs lfs) ?cpg) = c2t ?env ?cpg" + proof (rule lifts_c2t) + from perms_wf_cpg_test[OF assms(1), of sps] + show "wf_cpg_test (perm_ss sps sts) (perms (length sts) sps cpg) = (True, perm_ss sps sts')" . + next + show "length (perm_ss sps env) = length (perm_ss sps sts)" + by (metis assms(2) perm_ss_len) + qed + also have "\ = c2t env cpg" + proof(fold assms(2), rule perms_c2t) + from assms(1) show " wf_cpg_test sts cpg = (True, sts')" . + next + from assms(2) show "length env = length sts" . + qed + finally show ?thesis + by (unfold adjust_env_def adjust_cpg_def, simp) +qed + +lemma adjust_wf_cpg_test: + assumes "wf_cpg_test sts cpg = (True, sts')" + shows "wf_cpg_test (adjust_sts sps lfs sts) (adjust_cpg (length sts) sps (red_lfs lfs) cpg) = + (True, adjust_sts sps lfs sts')" +proof - + have " wf_cpg_test (lift_ss lfs (perm_ss sps sts)) (lift_ts (red_lfs lfs) (perms (length sts) sps cpg)) = + (True, lift_ss lfs (perm_ss sps sts'))" + proof(rule lifts_wf_cpg_test) + show " wf_cpg_test (perm_ss sps sts) (perms (length sts) sps cpg) = (True, perm_ss sps sts')" + by (rule perms_wf_cpg_test[OF assms]) + qed + thus ?thesis + by (unfold adjust_sts_def adjust_cpg_def, simp) +qed + +lemma sts_disj_test_correct: + assumes "sts_disj_test sts1 sts2" + shows "sts_disj sts1 sts2" + using assms +proof(induct sts1 arbitrary:sts2) + case (Nil sts2) + note Nil_1 = Nil + show ?case + proof(cases sts2) + case Nil + with Nil_1 + show ?thesis by (simp add:sts_disj_def) + next + case (Cons s2 ss2) + with Nil_1 show ?thesis by simp + qed +next + case (Cons s1 ss1 sts2) + note Cons_1 = Cons + show ?case + proof(cases "sts2") + case Nil + with Cons_1 show ?thesis by simp + next + case (Cons s2 ss2) + show ?thesis + proof(cases "s1 = Bound \ s2 = Bound") + case True + with Cons_1 Cons + show ?thesis by simp + next + case False + with Cons_1 Cons + have "sts_disj_test ss1 ss2" by (auto split:status.splits) + from Cons_1(1) [OF this] False + show ?thesis + apply (unfold Cons) + apply (unfold sts_disj_def) + by (smt False length_Suc_conv list.size(4) nth_Cons') + qed + qed +qed + +lemma sts_minus_free: + assumes "length sts1 = length sts2" + and "list_all (op = Free) sts2" + shows "sts1 - sts2 = sts1" + using assms +proof(induct sts1 arbitrary:sts2) + case (Nil sts2) + thus ?case by simp +next + case (Cons s1 ss1 sts2) + note Cons_1 = Cons + thus ?case + proof(cases sts2) + case Nil + with Cons + show ?thesis by simp + next + case (Cons s2 ss2) + have "ss1 - ss2 = ss1" + proof(rule Cons_1(1)) + show "length ss1 = length ss2" + by (metis Cons Cons_1(2) Suc_length_conv list.inject) + next + show "list_all (op = Free) ss2" + by (metis Cons Cons_1(3) list_all_simps(1)) + qed + moreover from Cons_1(3) Cons have "s2 = Free" + by (metis (full_types) list_all_simps(1)) + ultimately show ?thesis using Cons + apply simp + by (metis (hide_lams, mono_tags) minus_status.simps(2) minus_status.simps(3) status.exhaust) + qed +qed + +lemma st_simp [simp]: "St (nat_of x) = x" + by (metis nat_of.simps tstate.exhaust) + +lemma wf_cpg_test_len: + assumes "wf_cpg_test sts cpg = (b, sts')" + shows "length sts' = length sts" + using assms +proof(induct cpg arbitrary:sts sts' b) + case (CInstr instr sts sts' b) + then obtain a1 s1 a2 s2 where + eq_instr: "instr = ((a1, St s1), (a2, St s2))" + by (metis st_simp surj_pair) + with CInstr + show ?case by simp +qed (auto split:prod.splits) + +lemma wf_cpg_test_seq: + assumes "wf_cpg_test sts1 c1 = (True, sts1')" + and "wf_cpg_test sts2 c2 = (True, sts2')" + and "length sts1 = length sts2" + and "list_all (op = Free) sts1" + and "list_all (op = Free) sts2" + and "sts_disj_test sts1' sts2'" + shows "wf_cpg_test sts1 (CSeq c1 c2) = (True, sts1' + sts2')" +proof - + have "wf_cpg_test (sts1' + sts2) c2 = (True, sts1' + sts2')" + by (metis add_imp_eq assms(2) assms(5) assms(6) length_sts_plus + plus_list_len sts_disj_test_correct sts_minus_free wf_cpg_test_disj wf_cpg_test_extrapo wf_cpg_test_len) + hence "wf_cpg_test sts1' c2 = (True, sts1' + sts2')" + by (metis all_free_plus assms(1) assms(3) assms(5) wf_cpg_test_len) + with assms(1) + show ?thesis by simp +qed + +lemma c2t_seq: + assumes "c2t env c1 = t1" + and "c2t env c2 = t2" + shows "c2t env (CSeq c1 c2) = (t1; t2)" + using assms by simp + +lemma c2t_local: + assumes "\x. (c2t (x#xs) cpg = body x)" + shows "c2t xs (CLocal cpg) = (TL x. body x)" + using assms + by simp + +lemma wf_cpg_test_local: + assumes "wf_cpg_test (Free#sts) cpg = (b, s'#sts')" + shows "wf_cpg_test sts (CLocal cpg) = (b, sts')" + by (simp add:assms) + +lemma wf_c2t_combined: + assumes "wf_cpg_test sts cpg = (True, sts)" + and "c2t env cpg = tpg" + and "list_all (op = Free) sts" + and "length env = length sts" + shows "\ i. \ j s. ((i:[tpg]:j) s)" +proof + fix i + from wf_cpg_test_correct[OF assms(1), rule_format, of i] + obtain j where "c2p (sts - sts) i cpg j" by metis + from this[unfolded c2p_def] + obtain f where h: "\x. length x = length (sts - sts) \ + (\k x ! k = f i k) \ + Ex (i :[ c2t x cpg ]: j)" by metis + have "\ s. (i :[ c2t env cpg ]: j) s" + proof(rule h[rule_format], rule conjI) + show "length env = length (sts - sts)" + by (smt assms(4) minus_list_len) + next + show "\k env ! k = f i k" + by (metis assms(3) minus_status.simps(1) nth_sts_minus status.distinct(1) sts_minus_free) + qed + show "\ j s. ((i:[tpg]:j) s)" + by (metis `\s. (i :[ c2t env cpg ]: j) s` assms(2)) +qed + +subsection {* The Checker *} + +ML {* + print_depth 200 +*} + +subsubsection {* Auxilary functions *} + +ML {* +local + fun clear_binds ctxt = (ctxt |> Variable.binds_of |> Vartab.keys |> map (fn xi => (xi, NONE)) + |> fold Variable.bind_term) ctxt + fun get_binds ctxt = ctxt |> Variable.binds_of |> Vartab.dest |> map (fn (xi, (_, tm)) => (xi, SOME tm)) + fun set_binds blist ctxt = (fold Variable.bind_term blist) (clear_binds ctxt) +in + fun blocalM f = liftM (m2M (fn ctxt => returnM (get_binds ctxt))) + :|-- (fn binds => + f + :|-- (fn result => + liftM (m2M (fn ctxt' => s2M (set_binds binds ctxt') |-- returnM result + ))) + ) +end + + fun condM bf scan = (fn v => m0M (fn st => if (bf (v, st)) then scan v else returnM v)) + + local + val counter = Unsynchronized.ref 0 + in + fun init_counter n = (counter := n) + fun counter_test x = + if !counter <= 1 then true + else (counter := !counter - 1; false) + end + + (* break point monad *) + fun bpM v' = (fn v => m0M (fn st => raiseM (v', (v, st)))) + + fun the_theory () = ML_Context.the_local_context () |> Proof_Context.theory_of + fun the_context () = ML_Context.the_local_context () + + (* Calculating the numberal of integer [i] *) + fun nat_of i = if i = 0 then @{term "0::nat"} else + (Const ("Num.numeral_class.numeral", @{typ "num \ nat"}) $ + (Numeral.mk_cnumeral i |> term_of)) + + fun vfixM nm typ = (m2M' (fn ctxt => let + val ([x], ctxt') = Variable.variant_fixes [nm] ctxt + val tm_x = Free (x, typ) + in s2M ctxt' |-- returnM tm_x end)) + fun fixM nm typ = (m2M' (fn ctxt => let + val ([x], ctxt') = Variable.add_fixes [nm] ctxt + val tm_x = Free (x, typ) + in s2M ctxt' |-- returnM tm_x end)) + local + fun mk_listM l = + case l of + [] => @{fterm "[]"} + | (tm::tms) => localM (@{match "?x"} tm + |-- (mk_listM tms) + :|-- @{match "?xs"} + |-- @{fterm "?x#?xs"}) + in + fun mk_list_term ctxt l = [((), ctxt)] |> mk_listM l |> normVal |> fst + end + fun term_name (Const (x, _)) = Long_Name.base_name x + | term_name (Free (x, _)) = x + | term_name (Var ((x, _), _)) = x + | term_name _ = Name.uu; + + val rew_conv = Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)); + + fun simpl_conv ss thl ctm = + rew_conv (ss addsimps thl) ctm RS meta_eq_to_obj_eq; + + fun find_thms ctxt pats = + Find_Theorems.find_theorems ctxt NONE NONE true + (map (fn pat =>(true, Find_Theorems.Pattern + (Proof_Context.read_term_pattern ctxt pat))) pats) |> snd |> map snd + + + fun local_on arg rhs = [((), @{context})] |> + @{match "?body"} (Term.lambda arg rhs) |-- + @{fterm "TL x. ?body x"} |> normVal |> fst + fun find_idx vars l = (nat_of (find_index (equal l) vars)) + + local + fun mk_pair_term (i, j) = [((), @{context})] |> + @{match "?i"} (nat_of i) + |-- @{match "?j"} (nat_of j) + |-- @{fterm "(?i, ?j)"} |> normVal |> fst + in + fun mk_npair_list_term ctxt pair_list = + if pair_list = [] then @{term "[]::(nat \ nat) list"} + else pair_list |> map mk_pair_term |> mk_list_term ctxt + end + + fun list_of_array ary = let + val len = Array.length ary + val idx = upto (0, len - 1) + in map (fn i => Array.sub (ary, i)) idx end + +local + fun mk_env_term ctxt lst = + if lst = [] then @{term "[]::tstate list"} else (mk_list_term ctxt lst) + fun mk_pair_term ctxt (i, j) = [((), ctxt)] |> + @{match "?i"} (mk_env_term ctxt i) + |-- @{match "?j"} (nat_of j) + |-- @{fterm "(?i, ?j)"} |> normVal |> fst +in + fun mk_tpair_list_term ctxt pair_list = + if pair_list = [] then @{term "[] :: (tstate list \ nat) list"} + else pair_list |> map (mk_pair_term ctxt) |> mk_list_term ctxt +end + +*} + +subsubsection {* The reifier *} + +ML {* + fun locM (c2t_thm, test_thm) = (m1M' (fn env => + let + val Free (x, _) = hd env + val c2t_thm = Drule.generalize ([], [x]) c2t_thm + val c2t_thm = @{thm c2t_local} OF [c2t_thm] + val test_thm = @{thm wf_cpg_test_local} OF [test_thm] + in + s1M (tl env) |-- returnM (c2t_thm, test_thm) + end)) + + fun reify_local reify t = + ( @{match "TL x . ?body (x::tstate)"} t + |-- vfixM "x" @{typ "tstate"} + :|-- @{match "?x"} + :|-- (fn tmx => m1M' (fn env => s1M (tmx::env))) + |-- @{fterm "?body ?x"} + :|-- reify + :|-- locM + (* :|-- condM counter_test (bpM ("local", t)) *) + ) + + fun labelM exp = m0M' (fn (env, ctxt) => let + (* The following three lines are used for debugging purpose + (* (* The following two lines are used to set breakpoint counter + and invoke the reifyer in debug mode *) + val _ = init_counter 3 + val t = reify tpg_def_rhs [(rev tpg_def_largs_sufx, ctxt_tpg_def)] |> normExcp + *) + (* The following line is used to extract break point information and + establish the environment to execute body statements *) + val ((brc, exp), (_, (env, ctxt)::_)) = t + *) + val c2t_thm = [((), ctxt)] |> + @{match "?cpg"} exp + |-- @{match "?env"} (env |> mk_list_term ctxt) + |-- @{fterm "c2t ?env ?cpg"} |> normVal |> fst |> cterm_of (Proof_Context.theory_of ctxt) + |> simpl_conv (simpset_of ctxt) [] + val test_thm = [((), ctxt)] |> + @{match "?cpg"} exp + |-- @{match "?sts"} (env |> map (fn _ => @{term "Free"}) |> mk_list_term ctxt) + |-- @{fterm "wf_cpg_test ?sts ?cpg"} |> normVal |> fst |> cterm_of (Proof_Context.theory_of ctxt) + |> simpl_conv (simpset_of ctxt) [] + in returnM (c2t_thm, test_thm) end) + + fun reify_label t = + @{match "TLabel ?L"} t + |-- @{fterm "?L"} + :|-- (fn l => m1M' (fn st => returnM (find_idx st l))) + :|-- @{match ?L1} + |-- @{fterm "CLabel ?L1"} + (* :|-- condM counter_test (bpM ("label", t)) *) + :|-- labelM + + fun seqM ((c2t_thm1, test_thm1), (c2t_thm2, test_thm2)) = + m0M' (fn (env, ctxt) => + let + val simp_trans = (simpset_of ctxt) delsimps @{thms wf_cpg_test.simps c2t.simps} |> full_simplify + val ct2_thm = (@{thm c2t_seq} OF [c2t_thm1, c2t_thm2]) |> simp_trans + val test_thm = (@{thm wf_cpg_test_seq} OF [test_thm1, test_thm2]) |> simp_trans + in returnM (ct2_thm, test_thm) end) + + fun reify_seq reify t = + @{match "?c1; ?c2"} t + |-- ((@{fterm "?c1"} :|-- reify) -- + (@{fterm "?c2"} :|-- reify)) + (* :|-- condM counter_test (bpM ("seq", t)) *) + :|-- seqM + + fun reify_instr t = + @{match "\ ((?A0, ?L0), (?A1, ?L1))"} t + |-- @{fterm "?L0"} + :|-- (fn l => m1M' (fn st => returnM (find_idx st l))) + :|-- @{match ?L0'} + |-- @{fterm "?L1"} + :|-- (fn l => m1M' (fn st => returnM (find_idx st l))) + :|-- @{match ?L1'} + |-- @{fterm "CInstr ((?A0, ?L0'), (?A1, St ?L1'))"} + :|-- labelM + (* :|-- condM counter_test (bpM ("instr", t)) *) + + fun reify_var var = + (* condM counter_test (bpM ("var", var)) () |-- *) + (m0M' (fn (env, ctxt) => let + (* The following three lines are used for debugging purpose + (* (* The following two lines are used to set breakpoint counter + and invoke the reifyer in debug mode *) + val _ = init_counter 3 + val t = reify tpg_def_rhs [(rev tpg_def_largs_sufx, ctxt_tpg_def)] |> normExcp + *) + (* The following line is used to extract break point information and + establish the environment to execute body statements *) + val ((brc, var), (_, (env, ctxt)::_)) = t + *) + val (var_hd, var_args) = Term.strip_comb var + val (var_args_prefx, var_args_sufx) = + take_suffix (fn tm => type_of tm = @{typ "tstate"}) var_args + val var_skel_hd_typ = var_args_prefx |> map type_of |> (fn typs => typs ---> @{typ "cpg"}) + (* We discriminate two cases, one for tpg constants; the other for argument variable *) + val ([var_skel_hd_name], ctxt1) = + case var_hd of + (Const (nm, _)) => ([((nm |> Long_Name.base_name)^"_skel")], ctxt) + | _ => Variable.variant_fixes [(term_name var_hd^"_skel_")] ctxt + (* If [var_hd] is a constant, a corresponding skeleton constant is assumed to exist alrady *) + val var_skel_hd = if (Term.is_Const var_hd) then Syntax.read_term ctxt1 var_skel_hd_name + else Free (var_skel_hd_name, var_skel_hd_typ) + (* [skel_tm] is the skeleton object the properties of which will either be assumed (in case of + argument variable), or proved (in case of global constants ) *) + val skel_tm = Term.list_comb (var_skel_hd, var_args_prefx) + (* Start to prove or assume [c2t] property (named [c2t_thm]) of the skeleton object, + since the [c2t] property needs to be universally qantified, we + need to invent quantifier names: *) + val (var_skel_args_sufx_names, ctxt2) = + Variable.variant_fixes (var_args_sufx |> map term_name) ctxt1 + val var_skel_args_sufx = var_skel_args_sufx_names |> map (fn nm => Free (nm, @{typ "tstate"})) + val c2t_rhs = Term.list_comb (var_hd, var_args_prefx@var_skel_args_sufx) + val c2t_env = mk_list_term ctxt2 (var_skel_args_sufx |> rev) + val eqn = [((), ctxt2)] |> + @{match ?env} c2t_env + |-- @{match ?skel_tm} skel_tm + |-- @{match ?c2t_rhs} c2t_rhs + |-- @{fterm "Trueprop (c2t ?env ?skel_tm = ?c2t_rhs)"} |> normVal |> fst + fun all_on ctxt arg body = Const ("all", dummyT) $ (Term.lambda arg body) |> + Syntax.check_term ctxt + val c2t_eqn = fold (all_on ctxt2) (rev var_skel_args_sufx) eqn |> cterm_of (Proof_Context.theory_of ctxt2) + val ([c2t_thm], ctxt3) = + if (Term.is_Const var_hd) then + (* if [var_hd] is an constant, try to prove [c2t_eqn] by searching + into the facts database *) + let + val pat_skel_args = fold (curry (op ^)) (map (K " _ ") var_args_prefx) "" + val pat_skel_str = "( "^ var_skel_hd_name ^ pat_skel_args ^" )" + val test_pat = ("wf_cpg_test _ "^ pat_skel_str ^" = (True, _)") + val c2t_pat = ("c2t _ "^ pat_skel_str ^" = _") + val (wf_test_thms, c2t_thms) = ([test_pat], [c2t_pat]) |> pairself (find_thms ctxt2) + in + ([([((0, @{thm "refl"}), ctxt2)] |> + goalM (c2t_eqn |> term_of) + |-- tacM (fn ctxt => (Simplifier.simp_tac ((simpset_of ctxt) addsimps c2t_thms) 1)) + >> Goal.conclude |> normVal |> fst |> Raw_Simplifier.norm_hhf)], ctxt2) + end + else (* Otherwise, make [c2t_eqn] an assumption *) + Assumption.add_assumes [c2t_eqn] ctxt2 + (* Start to prove or assume [wf_cpg_test] property (named [wf_test_thm]) of the skeleton object. *) + val sts = map (fn tm => @{term "Free"}) var_args_sufx |> mk_list_term ctxt3 + val wf_test_eqn = [((), ctxt3)] |> + @{match ?cpg} skel_tm + |-- @{match ?sts} sts + |-- @{fterm "Trueprop (wf_cpg_test ?sts ?cpg = (True, ?sts))"} |> normVal |> fst + |> cterm_of (Proof_Context.theory_of ctxt3) + val ([wf_test_thm], ctxt4) = + if (Term.is_Const var_hd) then + let + val pat_skel_args = fold (curry (op ^)) (map (K " _ ") var_args_prefx) "" + val pat_skel_str = "( "^ var_skel_hd_name ^ pat_skel_args ^ " )" + val test_pat = ("wf_cpg_test _ "^ pat_skel_str ^" = (True, _)") + val c2t_pat = ("c2t _ "^ pat_skel_str ^" = _") + val wf_test_thms = [test_pat] |> (find_thms ctxt2) + in + ([([((0, @{thm "refl"}), ctxt2)] |> + goalM (wf_test_eqn |> term_of) + |-- tacM (fn ctxt => (Simplifier.simp_tac ((simpset_of ctxt) addsimps wf_test_thms) 1)) + >> Goal.conclude |> normVal |> fst |> Raw_Simplifier.norm_hhf)], ctxt3) + end + else Assumption.add_assumes [wf_test_eqn] ctxt3 + (* Start the derivation of the length theorem *) + val length_env = mk_list_term ctxt4 (var_args_sufx |> rev) + val length_thm = [((0, @{thm "init"}), ctxt4)] |> + @{match "(?env)"} length_env + |-- @{match "(?sts)"} sts + |-- @{fterm "Trueprop (length (?env::tstate list) = length (?sts::status list))"} + :|-- goalM + |-- tacM (fn ctxt => (Simplifier.simp_tac (simpset_of ctxt) 1)) + >> Goal.conclude |> normVal |> fst + (* Start compute two adjust operations, namely [sps] and [lfs] *) + val locs = var_args_sufx |> map (fn arg => find_index (equal arg) env) |> rev + val swaps = swaps_of locs + val sps = swaps |> mk_npair_list_term @{context} + val locs' = sexec swaps (Array.fromList locs) |> list_of_array + val pairs = ((~1::locs') ~~ (locs' @ [length env])) + fun lfs_of (t, ops) [] = ops |> rev + | lfs_of (t, ops) ((i, j)::pairs) = let + val stuf = upto (i + 1, j - 1) |> map (fn idx => nth env idx) + in if (stuf <> []) then lfs_of (t + length stuf + 1, (stuf, t)::ops) pairs + else lfs_of (t + length stuf + 1, ops) pairs + end + val lfs = lfs_of (0, []) pairs |> mk_tpair_list_term @{context} + (* [simp_trans] is the simplification procedure used to simply the theorem after + instantiation. + *) + val simp_trans = full_simplify ((simpset_of @{context}) addsimps @{thms adjust_sts_def + adjust_env_def perm_s_def perm_b_def map_idx_len + map_idx_def upto_map upto_empty} @ [c2t_thm]) + (* Instantiating adjust theorems *) + val adjust_c2t_thm = [((), ctxt4)] |> + @{match "?sps"} sps + |-- @{match "?lfs"} lfs + |-- thm_instM (@{thm adjust_c2t} OF [wf_test_thm, length_thm]) + |> normVal |> fst |> simp_trans + val adjust_test_thm = [((), ctxt4)] |> + @{match "?sps"} sps + |-- @{match "?lfs"} lfs + |-- thm_instM (@{thm adjust_wf_cpg_test} OF [wf_test_thm]) + |> normVal |> fst |> simp_trans + in + (* s2M ctxt4 |-- *) returnM (adjust_c2t_thm, adjust_test_thm) +end)) + + fun reify t = + localM (reify_seq reify t || + reify_local reify t || + reify_label t || + reify_instr t || + reify_var t + ) +*} + +subsubsection {* The Checker packed up as the asmb attribute *} + +ML {* + fun asmb_attrib def_thm = + Context.cases (fn thy => + (* val thy = @{theory} *) let + fun thy_exit ctxt = + Context.Theory (Local_Theory.exit_global (Local_Theory.assert_bottom true ctxt)) + val ctxt0 = Named_Target.theory_init thy + val (((x, y), [tpg_def]), ctxt_tpg_def) = Variable.import true [def_thm] ctxt0 + val (tpg_def_lhs, tpg_def_rhs) = [((), ctxt_tpg_def)] |> + @{match "Trueprop (?L = ?R)"} (prop_of tpg_def) + |-- @{fterm "?L"} -- @{fterm "?R"} |> normVal |> fst + val (tpg_def_lhd, tpg_def_largs) = Term.strip_comb tpg_def_lhs + val (tpg_def_largs_prefx, tpg_def_largs_sufx) = + take_suffix (fn tm => type_of tm = @{typ "tstate"}) tpg_def_largs + (* Invoking the reifyer in normal mode *) + val ((c2t_thm_1, test_thm_1), ((_, ctxt_r)::y)) = + reify tpg_def_rhs [(rev tpg_def_largs_sufx, ctxt_tpg_def)] + |> normVal + val asmb_thm_1 = (@{thm wf_c2t_combined} OF [test_thm_1, c2t_thm_1]) |> (full_simplify (simpset_of ctxt_r)) + val (r_cpg, r_tpg) = [((), ctxt_r)] |> + @{match "Trueprop (c2t _ ?X = ?tpg)"} (c2t_thm_1 |> prop_of) + |-- (@{fterm "?X"} -- @{fterm "?tpg"}) |> normVal |> fst + val tpg_def_params = Variable.add_fixed ctxt_tpg_def (tpg_def_lhs) [] |> map fst + |> sort (Variable.fixed_ord ctxt_tpg_def) + val r_cpg_frees = Term.add_frees r_cpg [] + local fun condense [] = [] + | condense xs = [hd xs] + in + val skel_def_params = + tpg_def_params |> map (fn nm => condense + (filter (fn (tnm, _) => String.isPrefix nm tnm) r_cpg_frees)) + |> flat |> map Free + end + val skel_def_rhs = fold Term.lambda (skel_def_params |> rev) r_cpg + local + val Const (nm, _) = tpg_def_lhs |> Term.head_of + in + val tpg_def_name = nm |> Long_Name.base_name + val skel_def_lhs = Free (tpg_def_name^"_skel", type_of skel_def_rhs) + end + val skel_def_eqn = [((), ctxt_r)] |> + @{match "?lhs"} skel_def_lhs + |-- @{match "?rhs"} skel_def_rhs + |-- @{fterm "Trueprop (?lhs = ?rhs)"} |> normVal |> fst + val ((skel_def_lhs, (skel_def_name, skel_def_thm)), lthy2) = + Specification.definition (NONE, (Attrib.empty_binding, skel_def_eqn)) ctxt_r + val c2t_thm_final = [((0, @{thm refl}), lthy2)] |> + @{match "?env"} (mk_list_term lthy2 (rev tpg_def_largs_sufx)) + |-- @{match "?skel"} (Term.list_comb (skel_def_lhs, skel_def_params)) + val c2t_thm_final = [((0, @{thm refl}), lthy2)] |> + @{match "?env"} (mk_list_term lthy2 (rev tpg_def_largs_sufx)) + |-- @{match "?skel"} (Term.list_comb (skel_def_lhs, skel_def_params)) + |-- @{match "?tpg"} tpg_def_lhs + |-- @{fterm "Trueprop (c2t ?env ?skel = ?tpg)"} + :|-- goalM + |-- tacM (fn ctxt => (Simplifier.simp_tac ((simpset_of ctxt |> Simplifier.clear_ss) + addsimps [skel_def_thm, c2t_thm_1]) 1)) + |-- tacM (fn ctxt => (Simplifier.simp_tac ((simpset_of ctxt |> Simplifier.clear_ss) + addsimps [def_thm]) 1)) + >> Goal.conclude |> normVal |> fst + val test_thm_final = [((0, @{thm refl}), lthy2)] |> + @{match "?sts"} (tpg_def_largs_sufx |> map (fn _ => @{term "Free"}) |> mk_list_term lthy2) + |-- @{match "?skel"} (Term.list_comb (skel_def_lhs, skel_def_params)) + |-- @{fterm "Trueprop (wf_cpg_test ?sts ?skel = (True, ?sts))"} + :|-- goalM + |-- tacM (fn ctxt => (Simplifier.simp_tac ((simpset_of ctxt |> Simplifier.clear_ss) + addsimps [skel_def_thm, test_thm_1]) 1)) + >> Goal.conclude |> normVal |> fst + val asmb_thm_final = [((0, @{thm refl}), lthy2)] |> + @{match "?tpg"} tpg_def_lhs + |-- @{fterm "Trueprop (\ i. \ j s. (i:[?tpg]:j) s)"} + :|-- goalM + |-- tacM (fn ctxt => (Simplifier.simp_tac ((simpset_of ctxt) + addsimps [tpg_def, asmb_thm_1]) 1)) + >> Goal.conclude |> normVal |> fst + fun generalize thm = let + val hyps = (#hyps (thm |> Thm.crep_thm)) + val thm' = if (length hyps = 0) then thm + else (fold Thm.implies_intr (#hyps (thm |> Thm.crep_thm) |> rev |> tl |> rev) thm) + in + thm' |> Thm.forall_intr_frees + end + val lthy3 = + Local_Theory.note ((Binding.name ("c2t_" ^ tpg_def_name ^ "_skel"), []), + [c2t_thm_final |> generalize]) lthy2 |> snd + val lthy4 = + Local_Theory.note ((Binding.name ("wf_" ^ tpg_def_name ^ "_skel"), []), + [test_thm_final |> generalize]) lthy3 |> snd + val lthy5 = + Local_Theory.note ((Binding.name ("asmb_" ^ tpg_def_name), []), + [asmb_thm_final |> Drule.export_without_context]) lthy4 |> snd +in + thy_exit lthy5 +end) (fn ctxt => Context.Proof ctxt) +*} + +setup {* + Attrib.setup @{binding asmb} (Scan.succeed (Thm.declaration_attribute asmb_attrib)) "asmb attribute" +*} + + +section {* Basic macros for TM *} + +definition [asmb]: "write_zero = (TL exit. \((W0, exit), (W0, exit)); TLabel exit)" + +definition [asmb]: "write_one = (TL exit. \((W1, exit), (W1, exit)); TLabel exit)" + +definition [asmb]: "move_left = (TL exit . \((L, exit), (L, exit)); TLabel exit)" + +definition [asmb]: "move_right = (TL exit . \((R, exit), (R, exit)); TLabel exit)" + +definition [asmb]: "if_one e = (TL exit . \((W0, exit), (W1, e)); TLabel exit)" + +definition [asmb]: "if_zero e = (TL exit . \((W0, e), (W1, exit)); TLabel exit)" + +definition [asmb]: "jmp e = \((W0, e), (W1, e))" + +definition [asmb]: + "right_until_zero = + (TL start exit. + TLabel start; + if_zero exit; + move_right; + jmp start; + TLabel exit + )" + +definition [asmb]: + "left_until_zero = + (TL start exit. + TLabel start; + if_zero exit; + move_left; + jmp start; + TLabel exit + )" + +definition [asmb]: + "right_until_one = + (TL start exit. + TLabel start; + if_one exit; + move_right; + jmp start; + TLabel exit + )" + +definition [asmb]: + "left_until_one = + (TL start exit. + TLabel start; + if_one exit; + move_left; + jmp start; + TLabel exit + )" + +definition [asmb]: + "left_until_double_zero = + (TL start exit. + TLabel start; + if_zero exit; + left_until_zero; + move_left; + if_one start; + TLabel exit)" + +definition [asmb]: + "shift_right = + (TL start exit. + TLabel start; + if_zero exit; + write_zero; + move_right; + right_until_zero; + write_one; + move_right; + jmp start; + TLabel exit + )" + +definition [asmb]: + "clear_until_zero = + (TL start exit. + TLabel start; + if_zero exit; + write_zero; + move_right; + jmp start; + TLabel exit)" + +definition [asmb]: + "shift_left = + (TL start exit. + TLabel start; + if_zero exit; + move_left; + write_one; + right_until_zero; + move_left; + write_zero; + move_right; + move_right; + jmp start; + TLabel exit) + " + +definition [asmb]: + "bone c1 c2 = (TL exit l_one. + if_one l_one; + (c1; + jmp exit); + TLabel l_one; + c2; + TLabel exit + )" + +definition [asmb]: + "cfill_until_one = (TL start exit. + TLabel start; + if_one exit; + write_one; + move_left; + jmp start; + TLabel exit + )" + +definition [asmb]: + "cmove = (TL start exit. + TLabel start; + left_until_zero; + left_until_one; + move_left; + if_zero exit; + move_right; + write_zero; + right_until_one; + right_until_zero; + write_one; + jmp start; + TLabel exit + )" + +definition [asmb]: + "cinit = (right_until_zero; move_right; write_one)" + +definition [asmb]: + "copy = (cinit; cmove; move_right; move_right; right_until_one; + move_left; move_left; cfill_until_one)" + +definition + "bzero c1 c2 = (TL exit l_zero. + if_zero l_zero; + (c1; + jmp exit); + TLabel l_zero; + c2; + TLabel exit + )" + +definition "if_reps_nz e = (move_right; + bzero (move_left; jmp e) (move_left) + )" + +declare if_reps_nz_def[unfolded bzero_def, asmb] + +definition "if_reps_z e = (move_right; + bone (move_left; jmp e) (move_left) + )" + +declare if_reps_z_def [unfolded bone_def, asmb] + +definition + "skip_or_set = bone (write_one; move_right; move_right) + (right_until_zero; move_right)" + +declare skip_or_set_def[unfolded bone_def, asmb] + +definition "tpg_fold tpgs = foldr TSeq (butlast tpgs) (last tpgs)" + +definition "cpg_fold cpgs = foldr CSeq (butlast cpgs) (last cpgs)" + +definition "skip_or_sets n = tpg_fold (replicate n skip_or_set)" + +definition "skip_or_sets_skel n = cpg_fold (replicate n skip_or_set_skel)" + +lemma c2t_skip_or_sets_skel: + "c2t [] (skip_or_sets_skel (Suc n)) = skip_or_sets (Suc n)" +proof(induct n) + case (Suc k) + thus ?case + apply (unfold skip_or_sets_skel_def cpg_fold_def skip_or_sets_def tpg_fold_def) + my_block + fix x k + have "(last (replicate (Suc k) x)) = x" + by (metis Suc_neq_Zero last_replicate) + my_block_end + apply (unfold this) + my_block + fix x k + have "(butlast (replicate (Suc k) x)) = replicate k x" + by (metis butlast_snoc replicate_Suc replicate_append_same) + my_block_end + apply (unfold this) + my_block + fix x k f y + have "foldr f (replicate (Suc k) x) y = f x (foldr f (replicate k x) y)" + by simp + my_block_end + apply (unfold this) + by (simp add:c2t_skip_or_set_skel) +next + case 0 + show ?case + by (simp add:skip_or_sets_skel_def cpg_fold_def skip_or_sets_def tpg_fold_def + c2t_skip_or_set_skel) +qed + +lemma wf_skip_or_sets_skel: + "wf_cpg_test [] (skip_or_sets_skel (Suc n)) = (True, [])" +proof(induct n) + case (Suc k) + thus ?case + apply (unfold skip_or_sets_skel_def cpg_fold_def) + my_block + fix x k + have "(last (replicate (Suc k) x)) = x" + by (metis Suc_neq_Zero last_replicate) + my_block_end + apply (unfold this) + my_block + fix x k + have "(butlast (replicate (Suc k) x)) = replicate k x" + by (metis butlast_snoc replicate_Suc replicate_append_same) + my_block_end + apply (unfold this) + my_block + fix x k f y + have "foldr f (replicate (Suc k) x) y = f x (foldr f (replicate k x) y)" + by simp + my_block_end + apply (unfold this) + by (simp add:wf_skip_or_set_skel) +next + case 0 + thus ?case + apply (unfold skip_or_sets_skel_def cpg_fold_def) + by (simp add:wf_skip_or_set_skel) +qed + +lemma asmb_skip_or_sets: + "\i. \j s. (i :[ skip_or_sets (Suc n) ]: j) s" + by (rule wf_c2t_combined[OF wf_skip_or_sets_skel c2t_skip_or_sets_skel], auto) + +definition [asmb]: "locate n = (skip_or_sets (Suc n); + move_left; + move_left; + left_until_zero; + move_right + )" + +definition [asmb]: "Inc a = locate a; + right_until_zero; + move_right; + shift_right; + move_left; + left_until_double_zero; + write_one; + left_until_double_zero; + move_right; + move_right + " + +definition [asmb]: "Dec a e = (TL continue. + (locate a; + if_reps_nz continue; + left_until_double_zero; + move_right; + move_right; + jmp e); + (TLabel continue; + right_until_zero; + move_left; + write_zero; + move_right; + move_right; + shift_left; + move_left; + move_left; + move_left; + left_until_double_zero; + move_right; + move_right))" + +end \ No newline at end of file diff -r 77daf1b85cf0 -r a5f5b9336007 thys2/MLs.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/MLs.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,61 @@ +theory MLs +imports Pure +keywords "ML_psf" :: prf_script and + "ML_goal" :: thy_goal +begin + +section {* ML for proof state transformation *} + +ML {* + structure ML_psf = Proof_Data ( + type T = Proof.state -> Proof.state + fun init _ = undefined + ) +*} + + +ML {* + fun ml_psf (txt, pos) state = + (((Proof.context_of state) |> + Context.proof_map (ML_Context.expression Position.none + "val psf : Proof.state -> Proof.state" + "Context.map_proof (ML_psf.put psf)" (ML_Lex.read pos txt)) |> + ML_psf.get) state):Proof.state +*} + +ML {* +val _ = + Outer_Syntax.command @{command_spec "ML_psf"} "ML state transfer function" + (Parse.ML_source >> (Toplevel.print oo (Toplevel.proof o ml_psf))) +*} + +section {* ML for global goal establishment *} + +ML {* + structure ML_goal = Proof_Data ( + type T = local_theory -> Proof.state + fun init _ = undefined + ) +*} + +ML {* + fun ml_goal (txt, pos) ctxt = + ((ctxt |> + Context.proof_map (ML_Context.expression Position.none + "val goalf : local_theory -> Proof.state" + "Context.map_proof (ML_goal.put goalf)" (ML_Lex.read pos txt)) |> + ML_goal.get) ctxt) :Proof.state +*} + +ML {* +val _ = + Outer_Syntax.local_theory_to_proof @{command_spec "ML_goal"} "ML goal function" + (Parse.ML_source >> ml_goal) +*} + +ML {* + fun start_theorem prop = Proof.theorem NONE (K I) [[(prop, [])]] +*} + +end + diff -r 77daf1b85cf0 -r a5f5b9336007 thys2/My_block.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/My_block.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,34 @@ +theory My_block +imports Main +keywords "my_block" :: prf_script and "my_block_end" :: prf_script + and "my_note" :: prf_decl +begin + +ML {* + val my_block = Proof.assert_backward #> Proof.enter_forward #> Proof.begin_block +*} + +ML {* + val my_block_end = Proof.end_block #> Proof.enter_backward +*} + +ML {* + val _ = Outer_Syntax.command @{command_spec "my_block"} "begin my block" + (Scan.succeed (Toplevel.print o (Toplevel.proof my_block))) + val _ = Outer_Syntax.command @{command_spec "my_block_end"} "begin my block" + (Scan.succeed (Toplevel.print o (Toplevel.proof my_block_end))) +*} + +ML {* + fun my_note x = + Proof.assert_backward #> Proof.enter_forward #> Proof.note_thmss_cmd x #> Proof.enter_backward +*} + +ML {* +val _ = + Outer_Syntax.command @{command_spec "my_note"} "define facts" + (Parse_Spec.name_facts >> (Toplevel.print oo (Toplevel.proof o my_note))); +*} + +end + diff -r 77daf1b85cf0 -r a5f5b9336007 thys2/ROOT --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/ROOT Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,20 @@ +session "Hoare_gen" = "HOL" + + options [document = pdf] + theories [document = false] + Hoare_gen + +session "Hoare_tm_basis" = "Hoare_gen" + + options [document = pdf] + theories [document = false, document_output = "output_tm", document_variants = "hoare_tm"] + Hoare_tm_basis + +session "Hoare_tm" = "Hoare_tm_basis" + + options [document = pdf] + theories [document = false, document_output = "output_tm", document_variants = "hoare_tm"] + Hoare_tm + +session "Hoare_abc" = "Hoare_tm" + + options [document = pdf] + theories [document = false, document_output = "./output_abc"] + Hoare_abc + diff -r 77daf1b85cf0 -r a5f5b9336007 thys2/ROOT~ --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/ROOT~ Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,24 @@ +session "Hoare_gen" = "HOL" + + options [document = pdf] + theories [document = false] + Hoare_gen + +session "Hoare_tm_basis" = "Hoare_gen" + + options [document = pdf] + theories [document = false, document_output = "output_tm", document_variants = "hoare_tm"] + Hoare_tm_basis + files + "document/root_tm.tex" + +session "Hoare_tm" = "Hoare_tm_basis" + + options [document = pdf] + theories [document = false, document_output = "output_tm", document_variants = "hoare_tm"] + Hoare_tm + files + "document/root_tm.tex" + +session "Hoare_abc" = "Hoare_tm" + + options [document = pdf] + theories [document = false, document_output = "./output_abc"] + Hoare_abc + diff -r 77daf1b85cf0 -r a5f5b9336007 thys2/Recs.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/Recs.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,939 @@ +theory Recs +imports Main Fact + "~~/src/HOL/Number_Theory/Primes" + "~~/src/HOL/Library/Nat_Bijection" + (* "~~/src/HOL/Library/Discrete" *) + Hoare_abc + LetElim +begin + +declare One_nat_def[simp del] + +(* + some definitions from + + A Course in Formal Languages, Automata and Groups + I M Chiswell + + and + + Lecture on undecidability + Michael M. Wolf +*) + +lemma if_zero_one [simp]: + "(if P then 1 else 0) = (0::nat) \ \ P" + "(0::nat) < (if P then 1 else 0) = P" + "(if P then 0 else 1) = (if \P then 1 else (0::nat))" +by (simp_all) + +lemma nth: + "(x # xs) ! 0 = x" + "(x # y # xs) ! 1 = y" + "(x # y # z # xs) ! 2 = z" + "(x # y # z # u # xs) ! 3 = u" +by (simp_all) + + +section {* Some auxiliary lemmas about @{text "\"} and @{text "\"} *} + +lemma setprod_atMost_Suc[simp]: + "(\i \ Suc n. f i) = (\i \ n. f i) * f(Suc n)" +by(simp add:atMost_Suc mult_ac) + +lemma setprod_lessThan_Suc[simp]: + "(\i < Suc n. f i) = (\i < n. f i) * f n" +by (simp add:lessThan_Suc mult_ac) + +lemma setsum_add_nat_ivl2: "n \ p \ + setsum f {.. nat" + shows "(\i < n. f i) = 0 \ (\i < n. f i = 0)" + "(\i \ n. f i) = 0 \ (\i \ n. f i = 0)" +by (auto) + +lemma setprod_eq_zero [simp]: + fixes f::"nat \ nat" + shows "(\i < n. f i) = 0 \ (\i < n. f i = 0)" + "(\i \ n. f i) = 0 \ (\i \ n. f i = 0)" +by (auto) + +lemma setsum_one_less: + fixes n::nat + assumes "\i < n. f i \ 1" + shows "(\i < n. f i) \ n" +using assms +by (induct n) (auto) + +lemma setsum_one_le: + fixes n::nat + assumes "\i \ n. f i \ 1" + shows "(\i \ n. f i) \ Suc n" +using assms +by (induct n) (auto) + +lemma setsum_eq_one_le: + fixes n::nat + assumes "\i \ n. f i = 1" + shows "(\i \ n. f i) = Suc n" +using assms +by (induct n) (auto) + +lemma setsum_least_eq: + fixes f::"nat \ nat" + assumes h0: "p \ n" + assumes h1: "\i \ {..i \ {p..n}. f i = 0" + shows "(\i \ n. f i) = p" +proof - + have eq_p: "(\i \ {..i \ {p..n}. f i) = 0" + using h2 by auto + have "(\i \ n. f i) = (\i \ {..i \ {p..n}. f i)" + using h0 by (simp add: setsum_add_nat_ivl2) + also have "... = (\i \ {..i \ n. f i) = p" using eq_p by simp +qed + +lemma nat_mult_le_one: + fixes m n::nat + assumes "m \ 1" "n \ 1" + shows "m * n \ 1" +using assms by (induct n) (auto) + +lemma setprod_one_le: + fixes f::"nat \ nat" + assumes "\i \ n. f i \ 1" + shows "(\i \ n. f i) \ 1" +using assms +by (induct n) (auto intro: nat_mult_le_one) + +lemma setprod_greater_zero: + fixes f::"nat \ nat" + assumes "\i \ n. f i \ 0" + shows "(\i \ n. f i) \ 0" +using assms by (induct n) (auto) + +lemma setprod_eq_one: + fixes f::"nat \ nat" + assumes "\i \ n. f i = Suc 0" + shows "(\i \ n. f i) = Suc 0" +using assms by (induct n) (auto) + +lemma setsum_cut_off_less: + fixes f::"nat \ nat" + assumes h1: "m \ n" + and h2: "\i \ {m..i < n. f i) = (\i < m. f i)" +proof - + have eq_zero: "(\i \ {m..i < n. f i) = (\i \ {..i \ {m..i \ {..i < n. f i) = (\i < m. f i)" by simp +qed + +lemma setsum_cut_off_le: + fixes f::"nat \ nat" + assumes h1: "m \ n" + and h2: "\i \ {m..n}. f i = 0" + shows "(\i \ n. f i) = (\i < m. f i)" +proof - + have eq_zero: "(\i \ {m..n}. f i) = 0" + using h2 by auto + have "(\i \ n. f i) = (\i \ {..i \ {m..n}. f i)" + using h1 by (simp add: setsum_add_nat_ivl2) + also have "... = (\i \ {..i \ n. f i) = (\i < m. f i)" by simp +qed + +lemma setprod_one [simp]: + fixes n::nat + shows "(\i < n. Suc 0) = Suc 0" + "(\i \ n. Suc 0) = Suc 0" +by (induct n) (simp_all) + +section {* Recursive Functions *} + +datatype recf = Z + | S + | Id nat nat + | Cn nat recf "recf list" + | Pr nat recf recf + | Mn nat recf + +fun arity :: "recf \ nat" + where + "arity Z = 1" +| "arity S = 1" +| "arity (Id m n) = m" +| "arity (Cn n f gs) = n" +| "arity (Pr n f g) = Suc n" +| "arity (Mn n f) = n" + +text {* Abbreviations for calculating the arity of the constructors *} + +abbreviation + "CN f gs \ Cn (arity (hd gs)) f gs" + +abbreviation + "PR f g \ Pr (arity f) f g" + +abbreviation + "MN f \ Mn (arity f - 1) f" + +text {* the evaluation function and termination relation *} + +fun rec_eval :: "recf \ nat list \ nat" + where + "rec_eval Z xs = 0" +| "rec_eval S xs = Suc (xs ! 0)" +| "rec_eval (Id m n) xs = xs ! n" +| "rec_eval (Cn n f gs) xs = rec_eval f (map (\x. rec_eval x xs) gs)" +| "rec_eval (Pr n f g) (0 # xs) = rec_eval f xs" +| "rec_eval (Pr n f g) (Suc x # xs) = + rec_eval g (x # (rec_eval (Pr n f g) (x # xs)) # xs)" +| "rec_eval (Mn n f) xs = (LEAST x. rec_eval f (x # xs) = 0)" + +inductive + terminates :: "recf \ nat list \ bool" +where + termi_z: "terminates Z [n]" +| termi_s: "terminates S [n]" +| termi_id: "\n < m; length xs = m\ \ terminates (Id m n) xs" +| termi_cn: "\terminates f (map (\g. rec_eval g xs) gs); + \g \ set gs. terminates g xs; length xs = n\ \ terminates (Cn n f gs) xs" +| termi_pr: "\\ y < x. terminates g (y # (rec_eval (Pr n f g) (y # xs) # xs)); + terminates f xs; + length xs = n\ + \ terminates (Pr n f g) (x # xs)" +| termi_mn: "\length xs = n; terminates f (r # xs); + rec_eval f (r # xs) = 0; + \ i < r. terminates f (i # xs) \ rec_eval f (i # xs) > 0\ \ terminates (Mn n f) xs" + + +section {* Arithmetic Functions *} + +text {* + @{text "constn n"} is the recursive function which computes + natural number @{text "n"}. +*} +fun constn :: "nat \ recf" + where + "constn 0 = Z" | + "constn (Suc n) = CN S [constn n]" + +definition + "rec_swap f = CN f [Id 2 1, Id 2 0]" + +definition + "rec_add = PR (Id 1 0) (CN S [Id 3 1])" + +definition + "rec_mult = PR Z (CN rec_add [Id 3 1, Id 3 2])" + +definition + "rec_power = rec_swap (PR (constn 1) (CN rec_mult [Id 3 1, Id 3 2]))" + +definition + "rec_fact_aux = PR (constn 1) (CN rec_mult [CN S [Id 3 0], Id 3 1])" + +definition + "rec_fact = CN rec_fact_aux [Id 1 0, Id 1 0]" + +definition + "rec_pred = CN (PR Z (Id 3 0)) [Id 1 0, Id 1 0]" + +definition + "rec_minus = rec_swap (PR (Id 1 0) (CN rec_pred [Id 3 1]))" + +lemma constn_lemma [simp]: + "rec_eval (constn n) xs = n" +by (induct n) (simp_all) + +lemma swap_lemma [simp]: + "rec_eval (rec_swap f) [x, y] = rec_eval f [y, x]" +by (simp add: rec_swap_def) + +lemma add_lemma [simp]: + "rec_eval rec_add [x, y] = x + y" +by (induct x) (simp_all add: rec_add_def) + +lemma mult_lemma [simp]: + "rec_eval rec_mult [x, y] = x * y" +by (induct x) (simp_all add: rec_mult_def) + +lemma power_lemma [simp]: + "rec_eval rec_power [x, y] = x ^ y" +by (induct y) (simp_all add: rec_power_def) + +lemma fact_aux_lemma [simp]: + "rec_eval rec_fact_aux [x, y] = fact x" +by (induct x) (simp_all add: rec_fact_aux_def) + +lemma fact_lemma [simp]: + "rec_eval rec_fact [x] = fact x" +by (simp add: rec_fact_def) + +lemma pred_lemma [simp]: + "rec_eval rec_pred [x] = x - 1" +by (induct x) (simp_all add: rec_pred_def) + +lemma minus_lemma [simp]: + "rec_eval rec_minus [x, y] = x - y" +by (induct y) (simp_all add: rec_minus_def) + + +section {* Logical functions *} + +text {* + The @{text "sign"} function returns 1 when the input argument + is greater than @{text "0"}. *} + +definition + "rec_sign = CN rec_minus [constn 1, CN rec_minus [constn 1, Id 1 0]]" + +definition + "rec_not = CN rec_minus [constn 1, Id 1 0]" + +text {* + @{text "rec_eq"} compares two arguments: returns @{text "1"} + if they are equal; @{text "0"} otherwise. *} +definition + "rec_eq = CN rec_minus [CN (constn 1) [Id 2 0], CN rec_add [rec_minus, rec_swap rec_minus]]" + +definition + "rec_noteq = CN rec_not [rec_eq]" + +definition + "rec_conj = CN rec_sign [rec_mult]" + +definition + "rec_disj = CN rec_sign [rec_add]" + +definition + "rec_imp = CN rec_disj [CN rec_not [Id 2 0], Id 2 1]" + +text {* @{term "rec_ifz [z, x, y]"} returns x if z is zero, + y otherwise; @{term "rec_if [z, x, y]"} returns x if z is *not* + zero, y otherwise *} + +definition + "rec_ifz = PR (Id 2 0) (Id 4 3)" + +definition + "rec_if = CN rec_ifz [CN rec_not [Id 3 0], Id 3 1, Id 3 2]" + + +lemma sign_lemma [simp]: + "rec_eval rec_sign [x] = (if x = 0 then 0 else 1)" +by (simp add: rec_sign_def) + +lemma not_lemma [simp]: + "rec_eval rec_not [x] = (if x = 0 then 1 else 0)" +by (simp add: rec_not_def) + +lemma eq_lemma [simp]: + "rec_eval rec_eq [x, y] = (if x = y then 1 else 0)" +by (simp add: rec_eq_def) + +lemma noteq_lemma [simp]: + "rec_eval rec_noteq [x, y] = (if x \ y then 1 else 0)" +by (simp add: rec_noteq_def) + +lemma conj_lemma [simp]: + "rec_eval rec_conj [x, y] = (if x = 0 \ y = 0 then 0 else 1)" +by (simp add: rec_conj_def) + +lemma disj_lemma [simp]: + "rec_eval rec_disj [x, y] = (if x = 0 \ y = 0 then 0 else 1)" +by (simp add: rec_disj_def) + +lemma imp_lemma [simp]: + "rec_eval rec_imp [x, y] = (if 0 < x \ y = 0 then 0 else 1)" +by (simp add: rec_imp_def) + +lemma ifz_lemma [simp]: + "rec_eval rec_ifz [z, x, y] = (if z = 0 then x else y)" +by (case_tac z) (simp_all add: rec_ifz_def) + +lemma if_lemma [simp]: + "rec_eval rec_if [z, x, y] = (if 0 < z then x else y)" +by (simp add: rec_if_def) + +section {* Less and Le Relations *} + +text {* + @{text "rec_less"} compares two arguments and returns @{text "1"} if + the first is less than the second; otherwise returns @{text "0"}. *} + +definition + "rec_less = CN rec_sign [rec_swap rec_minus]" + +definition + "rec_le = CN rec_disj [rec_less, rec_eq]" + +lemma less_lemma [simp]: + "rec_eval rec_less [x, y] = (if x < y then 1 else 0)" +by (simp add: rec_less_def) + +lemma le_lemma [simp]: + "rec_eval rec_le [x, y] = (if (x \ y) then 1 else 0)" +by(simp add: rec_le_def) + + +section {* Summation and Product Functions *} + +definition + "rec_sigma1 f = PR (CN f [CN Z [Id 1 0], Id 1 0]) + (CN rec_add [Id 3 1, CN f [CN S [Id 3 0], Id 3 2]])" + +definition + "rec_sigma2 f = PR (CN f [CN Z [Id 2 0], Id 2 0, Id 2 1]) + (CN rec_add [Id 4 1, CN f [CN S [Id 4 0], Id 4 2, Id 4 3]])" + +definition + "rec_accum1 f = PR (CN f [CN Z [Id 1 0], Id 1 0]) + (CN rec_mult [Id 3 1, CN f [CN S [Id 3 0], Id 3 2]])" + +definition + "rec_accum2 f = PR (CN f [CN Z [Id 2 0], Id 2 0, Id 2 1]) + (CN rec_mult [Id 4 1, CN f [CN S [Id 4 0], Id 4 2, Id 4 3]])" + +definition + "rec_accum3 f = PR (CN f [CN Z [Id 3 0], Id 3 0, Id 3 1, Id 3 2]) + (CN rec_mult [Id 5 1, CN f [CN S [Id 5 0], Id 5 2, Id 5 3, Id 5 4]])" + + +lemma sigma1_lemma [simp]: + shows "rec_eval (rec_sigma1 f) [x, y] = (\ z \ x. rec_eval f [z, y])" +by (induct x) (simp_all add: rec_sigma1_def) + +lemma sigma2_lemma [simp]: + shows "rec_eval (rec_sigma2 f) [x, y1, y2] = (\ z \ x. rec_eval f [z, y1, y2])" +by (induct x) (simp_all add: rec_sigma2_def) + +lemma accum1_lemma [simp]: + shows "rec_eval (rec_accum1 f) [x, y] = (\ z \ x. rec_eval f [z, y])" +by (induct x) (simp_all add: rec_accum1_def) + +lemma accum2_lemma [simp]: + shows "rec_eval (rec_accum2 f) [x, y1, y2] = (\ z \ x. rec_eval f [z, y1, y2])" +by (induct x) (simp_all add: rec_accum2_def) + +lemma accum3_lemma [simp]: + shows "rec_eval (rec_accum3 f) [x, y1, y2, y3] = (\ z \ x. (rec_eval f) [z, y1, y2, y3])" +by (induct x) (simp_all add: rec_accum3_def) + + +section {* Bounded Quantifiers *} + +definition + "rec_all1 f = CN rec_sign [rec_accum1 f]" + +definition + "rec_all2 f = CN rec_sign [rec_accum2 f]" + +definition + "rec_all3 f = CN rec_sign [rec_accum3 f]" + +definition + "rec_all1_less f = (let cond1 = CN rec_eq [Id 3 0, Id 3 1] in + let cond2 = CN f [Id 3 0, Id 3 2] + in CN (rec_all2 (CN rec_disj [cond1, cond2])) [Id 2 0, Id 2 0, Id 2 1])" + +definition + "rec_all2_less f = (let cond1 = CN rec_eq [Id 4 0, Id 4 1] in + let cond2 = CN f [Id 4 0, Id 4 2, Id 4 3] in + CN (rec_all3 (CN rec_disj [cond1, cond2])) [Id 3 0, Id 3 0, Id 3 1, Id 3 2])" + +definition + "rec_ex1 f = CN rec_sign [rec_sigma1 f]" + +definition + "rec_ex2 f = CN rec_sign [rec_sigma2 f]" + + +lemma ex1_lemma [simp]: + "rec_eval (rec_ex1 f) [x, y] = (if (\z \ x. 0 < rec_eval f [z, y]) then 1 else 0)" +by (simp add: rec_ex1_def) + +lemma ex2_lemma [simp]: + "rec_eval (rec_ex2 f) [x, y1, y2] = (if (\z \ x. 0 < rec_eval f [z, y1, y2]) then 1 else 0)" +by (simp add: rec_ex2_def) + +lemma all1_lemma [simp]: + "rec_eval (rec_all1 f) [x, y] = (if (\z \ x. 0 < rec_eval f [z, y]) then 1 else 0)" +by (simp add: rec_all1_def) + +lemma all2_lemma [simp]: + "rec_eval (rec_all2 f) [x, y1, y2] = (if (\z \ x. 0 < rec_eval f [z, y1, y2]) then 1 else 0)" +by (simp add: rec_all2_def) + +lemma all3_lemma [simp]: + "rec_eval (rec_all3 f) [x, y1, y2, y3] = (if (\z \ x. 0 < rec_eval f [z, y1, y2, y3]) then 1 else 0)" +by (simp add: rec_all3_def) + +lemma all1_less_lemma [simp]: + "rec_eval (rec_all1_less f) [x, y] = (if (\z < x. 0 < rec_eval f [z, y]) then 1 else 0)" +apply(auto simp add: Let_def rec_all1_less_def) +apply (metis nat_less_le)+ +done + +lemma all2_less_lemma [simp]: + "rec_eval (rec_all2_less f) [x, y1, y2] = (if (\z < x. 0 < rec_eval f [z, y1, y2]) then 1 else 0)" +apply(auto simp add: Let_def rec_all2_less_def) +apply(metis nat_less_le)+ +done + +section {* Quotients *} + +definition + "rec_quo = (let lhs = CN S [Id 3 0] in + let rhs = CN rec_mult [Id 3 2, CN S [Id 3 1]] in + let cond = CN rec_eq [lhs, rhs] in + let if_stmt = CN rec_if [cond, CN S [Id 3 1], Id 3 1] + in PR Z if_stmt)" + +fun Quo where + "Quo x 0 = 0" +| "Quo x (Suc y) = (if (Suc y = x * (Suc (Quo x y))) then Suc (Quo x y) else Quo x y)" + +lemma Quo0: + shows "Quo 0 y = 0" +by (induct y) (auto) + +lemma Quo1: + "x * (Quo x y) \ y" +by (induct y) (simp_all) + +lemma Quo2: + "b * (Quo b a) + a mod b = a" +by (induct a) (auto simp add: mod_Suc) + +lemma Quo3: + "n * (Quo n m) = m - m mod n" +using Quo2[of n m] by (auto) + +lemma Quo4: + assumes h: "0 < x" + shows "y < x + x * Quo x y" +proof - + have "x - (y mod x) > 0" using mod_less_divisor assms by auto + then have "y < y + (x - (y mod x))" by simp + then have "y < x + (y - (y mod x))" by simp + then show "y < x + x * (Quo x y)" by (simp add: Quo3) +qed + +lemma Quo_div: + shows "Quo x y = y div x" +apply(case_tac "x = 0") +apply(simp add: Quo0) +apply(subst split_div_lemma[symmetric]) +apply(auto intro: Quo1 Quo4) +done + +lemma Quo_rec_quo: + shows "rec_eval rec_quo [y, x] = Quo x y" +by (induct y) (simp_all add: rec_quo_def) + +lemma quo_lemma [simp]: + shows "rec_eval rec_quo [y, x] = y div x" +by (simp add: Quo_div Quo_rec_quo) + + +section {* Iteration *} + +definition + "rec_iter f = PR (Id 1 0) (CN f [Id 3 1])" + +fun Iter where + "Iter f 0 = id" +| "Iter f (Suc n) = f \ (Iter f n)" + +lemma Iter_comm: + "(Iter f n) (f x) = f ((Iter f n) x)" +by (induct n) (simp_all) + +lemma iter_lemma [simp]: + "rec_eval (rec_iter f) [n, x] = Iter (\x. rec_eval f [x]) n x" +by (induct n) (simp_all add: rec_iter_def) + + +section {* Bounded Maximisation *} + + +fun BMax_rec where + "BMax_rec RR 0 = 0" +| "BMax_rec RR (Suc n) = (if RR (Suc n) then (Suc n) else BMax_rec RR n)" + +definition + BMax_set :: "(nat \ bool) \ nat \ nat" +where + "BMax_set RR x = Max ({z. z \ x \ RR z} \ {0})" + +lemma BMax_rec_eq1: + "BMax_rec RR x = (GREATEST z. (RR z \ z \ x) \ z = 0)" +apply(induct x) +apply(auto intro: Greatest_equality Greatest_equality[symmetric]) +apply(simp add: le_Suc_eq) +by metis + +lemma BMax_rec_eq2: + "BMax_rec RR x = Max ({z. z \ x \ RR z} \ {0})" +apply(induct x) +apply(auto intro: Max_eqI Max_eqI[symmetric]) +apply(simp add: le_Suc_eq) +by metis + +lemma BMax_rec_eq3: + "BMax_rec RR x = Max (Set.filter (\z. RR z) {..x} \ {0})" +by (simp add: BMax_rec_eq2 Set.filter_def) + +definition + "rec_max1 f = PR Z (CN rec_ifz [CN f [CN S [Id 3 0], Id 3 2], CN S [Id 3 0], Id 3 1])" + +lemma max1_lemma [simp]: + "rec_eval (rec_max1 f) [x, y] = BMax_rec (\u. rec_eval f [u, y] = 0) x" +by (induct x) (simp_all add: rec_max1_def) + +definition + "rec_max2 f = PR Z (CN rec_ifz [CN f [CN S [Id 4 0], Id 4 2, Id 4 3], CN S [Id 4 0], Id 4 1])" + +lemma max2_lemma [simp]: + "rec_eval (rec_max2 f) [x, y1, y2] = BMax_rec (\u. rec_eval f [u, y1, y2] = 0) x" +by (induct x) (simp_all add: rec_max2_def) + + +section {* Encodings using Cantor's pairing function *} + +text {* + We use Cantor's pairing function from Nat_Bijection. + However, we need to prove that the formulation of the + decoding function there is recursive. For this we first + prove that we can extract the maximal triangle number + using @{term prod_decode}. +*} + +abbreviation Max_triangle_aux where + "Max_triangle_aux k z \ fst (prod_decode_aux k z) + snd (prod_decode_aux k z)" + +abbreviation Max_triangle where + "Max_triangle z \ Max_triangle_aux 0 z" + +abbreviation + "pdec1 z \ fst (prod_decode z)" + +abbreviation + "pdec2 z \ snd (prod_decode z)" + +abbreviation + "penc m n \ prod_encode (m, n)" + +lemma fst_prod_decode: + "pdec1 z = z - triangle (Max_triangle z)" +by (subst (3) prod_decode_inverse[symmetric]) + (simp add: prod_encode_def prod_decode_def split: prod.split) + +lemma snd_prod_decode: + "pdec2 z = Max_triangle z - pdec1 z" +by (simp only: prod_decode_def) + +lemma le_triangle: + "m \ triangle (n + m)" +by (induct_tac m) (simp_all) + +lemma Max_triangle_triangle_le: + "triangle (Max_triangle z) \ z" +by (subst (9) prod_decode_inverse[symmetric]) + (simp add: prod_decode_def prod_encode_def split: prod.split) + +lemma Max_triangle_le: + "Max_triangle z \ z" +proof - + have "Max_triangle z \ triangle (Max_triangle z)" + using le_triangle[of _ 0, simplified] by simp + also have "... \ z" by (rule Max_triangle_triangle_le) + finally show "Max_triangle z \ z" . +qed + +lemma w_aux: + "Max_triangle (triangle k + m) = Max_triangle_aux k m" +by (simp add: prod_decode_def[symmetric] prod_decode_triangle_add) + +lemma y_aux: "y \ Max_triangle_aux y k" +apply(induct k arbitrary: y rule: nat_less_induct) +apply(subst (1 2) prod_decode_aux.simps) +apply(simp) +apply(rule impI) +apply(drule_tac x="n - Suc y" in spec) +apply(drule mp) +apply(auto)[1] +apply(drule_tac x="Suc y" in spec) +apply(erule Suc_leD) +done + +lemma Max_triangle_greatest: + "Max_triangle z = (GREATEST k. (triangle k \ z \ k \ z) \ k = 0)" +apply(rule Greatest_equality[symmetric]) +apply(rule disjI1) +apply(rule conjI) +apply(rule Max_triangle_triangle_le) +apply(rule Max_triangle_le) +apply(erule disjE) +apply(erule conjE) +apply(subst (asm) (1) le_iff_add) +apply(erule exE) +apply(clarify) +apply(simp only: w_aux) +apply(rule y_aux) +apply(simp) +done + + +definition + "rec_triangle = CN rec_quo [CN rec_mult [Id 1 0, S], constn 2]" + +definition + "rec_max_triangle = + (let cond = CN rec_not [CN rec_le [CN rec_triangle [Id 2 0], Id 2 1]] in + CN (rec_max1 cond) [Id 1 0, Id 1 0])" + + +lemma triangle_lemma [simp]: + "rec_eval rec_triangle [x] = triangle x" +by (simp add: rec_triangle_def triangle_def) + +lemma max_triangle_lemma [simp]: + "rec_eval rec_max_triangle [x] = Max_triangle x" +by (simp add: Max_triangle_greatest rec_max_triangle_def Let_def BMax_rec_eq1) + + +text {* Encodings for Products *} + +definition + "rec_penc = CN rec_add [CN rec_triangle [CN rec_add [Id 2 0, Id 2 1]], Id 2 0]" + +definition + "rec_pdec1 = CN rec_minus [Id 1 0, CN rec_triangle [CN rec_max_triangle [Id 1 0]]]" + +definition + "rec_pdec2 = CN rec_minus [CN rec_max_triangle [Id 1 0], CN rec_pdec1 [Id 1 0]]" + +lemma pdec1_lemma [simp]: + "rec_eval rec_pdec1 [z] = pdec1 z" +by (simp add: rec_pdec1_def fst_prod_decode) + +lemma pdec2_lemma [simp]: + "rec_eval rec_pdec2 [z] = pdec2 z" +by (simp add: rec_pdec2_def snd_prod_decode) + +lemma penc_lemma [simp]: + "rec_eval rec_penc [m, n] = penc m n" +by (simp add: rec_penc_def prod_encode_def) + + +text {* Encodings of Lists *} + +fun + lenc :: "nat list \ nat" +where + "lenc [] = 0" +| "lenc (x # xs) = penc (Suc x) (lenc xs)" + +fun + ldec :: "nat \ nat \ nat" +where + "ldec z 0 = (pdec1 z) - 1" +| "ldec z (Suc n) = ldec (pdec2 z) n" + +lemma pdec_zero_simps [simp]: + "pdec1 0 = 0" + "pdec2 0 = 0" +by (simp_all add: prod_decode_def prod_decode_aux.simps) + +lemma ldec_zero: + "ldec 0 n = 0" +by (induct n) (simp_all add: prod_decode_def prod_decode_aux.simps) + +lemma list_encode_inverse: + "ldec (lenc xs) n = (if n < length xs then xs ! n else 0)" +by (induct xs arbitrary: n rule: lenc.induct) + (auto simp add: ldec_zero nth_Cons split: nat.splits) + +lemma lenc_length_le: + "length xs \ lenc xs" +by (induct xs) (simp_all add: prod_encode_def) + + +text {* Membership for the List Encoding *} + +fun within :: "nat \ nat \ bool" where + "within z 0 = (0 < z)" +| "within z (Suc n) = within (pdec2 z) n" + +definition enclen :: "nat \ nat" where + "enclen z = BMax_rec (\x. within z (x - 1)) z" + +lemma within_False [simp]: + "within 0 n = False" +by (induct n) (simp_all) + +lemma within_length [simp]: + "within (lenc xs) s = (s < length xs)" +apply(induct s arbitrary: xs) +apply(case_tac xs) +apply(simp_all add: prod_encode_def) +apply(case_tac xs) +apply(simp_all) +done + +text {* Length of Encoded Lists *} + +lemma enclen_length [simp]: + "enclen (lenc xs) = length xs" +unfolding enclen_def +apply(simp add: BMax_rec_eq1) +apply(rule Greatest_equality) +apply(auto simp add: lenc_length_le) +done + +lemma enclen_penc [simp]: + "enclen (penc (Suc x) (lenc xs)) = Suc (enclen (lenc xs))" +by (simp only: lenc.simps[symmetric] enclen_length) (simp) + +lemma enclen_zero [simp]: + "enclen 0 = 0" +by (simp add: enclen_def) + + +text {* Recursive Definitions for List Encodings *} + +fun + rec_lenc :: "recf list \ recf" +where + "rec_lenc [] = Z" +| "rec_lenc (f # fs) = CN rec_penc [CN S [f], rec_lenc fs]" + +definition + "rec_ldec = CN rec_pred [CN rec_pdec1 [rec_swap (rec_iter rec_pdec2)]]" + +definition + "rec_within = CN rec_less [Z, rec_swap (rec_iter rec_pdec2)]" + +definition + "rec_enclen = CN (rec_max1 (CN rec_not [CN rec_within [Id 2 1, CN rec_pred [Id 2 0]]])) [Id 1 0, Id 1 0]" + +lemma ldec_iter: + "ldec z n = pdec1 (Iter pdec2 n z) - 1" +by (induct n arbitrary: z) (simp | subst Iter_comm)+ + +lemma within_iter: + "within z n = (0 < Iter pdec2 n z)" +by (induct n arbitrary: z) (simp | subst Iter_comm)+ + +lemma lenc_lemma [simp]: + "rec_eval (rec_lenc fs) xs = lenc (map (\f. rec_eval f xs) fs)" +by (induct fs) (simp_all) + +lemma ldec_lemma [simp]: + "rec_eval rec_ldec [z, n] = ldec z n" +by (simp add: ldec_iter rec_ldec_def) + +lemma within_lemma [simp]: + "rec_eval rec_within [z, n] = (if within z n then 1 else 0)" +by (simp add: within_iter rec_within_def) + +lemma enclen_lemma [simp]: + "rec_eval rec_enclen [z] = enclen z" +by (simp add: rec_enclen_def enclen_def) + +section {* Compliation from Recs to Abacus *} + +definition "ms a xs = fam_conj {0..<(length xs)} (\ i. mm (a+i) (xs!i))" + +definition "za a = mm a 0" + +definition "zs I = fam_conj I za" + +consts mn_compute_g :: "(nat \ tpg) \ tpg" + mn_save_arg :: "nat \ tpg" + mn_restore_arg :: "nat \ tpg" + mn_install_f_arg :: "nat \ tpg" + +consts rec_ci :: "recf \ (tpg \ nat)" + +lemma mn_ci: + "rec_ci (Cn n f gs) = + (let (cf, muf) = rec_ci f in + let cgs = map (fst o rec_ci) gs in + let mgs = map (snd o rec_ci) gs in + let mub = int (Max (set (muf#mgs))) in + let ps = map nat [mub+1..mub + int (length gs)] in + let qs = map nat [mub + int (length gs) + 1 .. mub + int (n + length gs)] in + let compute_gs = tpg_fold (map mn_compute_g (zip ps cgs)) in + let save_args = tpg_fold (map mn_save_arg qs) in + let restore_args = tpg_fold (map mn_restore_arg qs) in + let clear_all = (tpg_fold (map (clear o nat) [0..(int (length gs))])) in + let install_f_args = tpg_fold (map mn_install_f_arg ps) in + let cf = (compute_gs; save_args; clear_all; install_f_args; cf; restore_args) in + let mu = ((nat mub) + n + length gs) in + (cf, mu))" + sorry + +lemma rec_ci_cnE: + assumes "P (rec_ci (Cn n f gs))" + obtains cf muf cgs mgs mub pss qs compute_gs + save_args restore_args clear_all install_f_args code mu + where "(cf, muf) = rec_ci f" + "cgs = map (fst o rec_ci) gs" + "mgs = map (snd o rec_ci) gs" + "mub = int (Max (set (muf#mgs)))" + "pss = map nat [mub+1..mub + int (length gs)]" + "qs = map nat [mub + int (length gs) + 1 .. mub + int (n + length gs)]" + "compute_gs = tpg_fold (map mn_compute_g (zip pss cgs))" + "save_args = tpg_fold (map mn_save_arg qs)" + "restore_args = tpg_fold (map mn_restore_arg qs)" + "clear_all = (tpg_fold (map (clear o nat) [0..(int (length gs))]))" + "install_f_args = tpg_fold (map mn_install_f_arg pss)" + "code = (compute_gs; save_args; clear_all; install_f_args; cf; restore_args)" + "mu = (nat mub) + n + length gs" + "P (code, mu)" +proof - + show ?thesis + apply (insert assms) + apply (unfold mn_ci) + apply (tactic {* let_elim_tac @{context} 1 *}) + apply (insert that) + by metis +qed + +theorem rec_ci_correct: + assumes "terminates rcf xs" + "rec_ci rcf = (c, mu)" + shows + "IA. \ pc i \* ms 0 xs \* mm (length xs) 0 \* zs {length xs<..(mu - length xs)}\ + i:[c]:j + \ pc j \* ms 0 xs \* mm (length xs) (rec_eval rcf xs) \* zs {length xs<..(mu - length xs)}\" + using assms +proof(induct arbitrary:c mu) + case (termi_cn f xs gs n c mu) + hence [unfolded mn_ci]: "rec_ci (Cn n f gs) = (c, mu)" by simp + thus ?case + proof(let_elim) + case (1 cf muf cgs mgs mub ps qs compute_gs + save_args restore_args clear_all install_f_args code) + from 1(1) have "c = code" "mu = (nat mub + n + length gs)" by auto + show ?thesis + proof(unfold `c = code`) + oops + +end + diff -r 77daf1b85cf0 -r a5f5b9336007 thys2/Sort_ops.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/Sort_ops.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,65 @@ +theory Sort_ops +imports Term_pat +begin + +ML {* +local + open Array +in + + fun swap i j a = let val ai = sub(a, i); + val _ = update(a, i, sub(a, j)) + val _ = update(a, j, ai) + in + a + end + + fun pre_ops_of a = + getM :|-- (fn (l, r, piv, i, j) => let + (* val _ = Output.tracing ("("^string_of_int i^","^string_of_int j^")") *) + in + if (j < i) then (returnM (swap l j a) |-- returnM [(l, j)]) + else (if (sub(a, i) <= piv andalso i <= r) then (setM (l, r, piv, i + 1, j) + |-- pre_ops_of a) + else if (sub(a, j) > piv ) then (setM (l, r, piv, i, j - 1) |-- pre_ops_of a) + else (pre_ops_of (swap i j a) :|-- (fn ops => returnM ((i,j)::ops)))) + end + ) + + fun ops_of i j a = + if (j < i) then [] + else let + val piv = sub(a, i) + val (ops1, (_, _, _, i', j')) = pre_ops_of a (i, j, piv, i, j) |> normVal + val ops2 = ops_of i (j' - 1) a + val ops3 = ops_of (j' + 1) j a + in + ops1 @ ops2 @ ops3 + end + + fun rem_sdup [] = [] + | rem_sdup [c] = [c] + | rem_sdup ((i, j)::(i', j')::xs) = if ((i = i' andalso j = j') orelse + (i = j' andalso j = i')) then rem_sdup (xs) + else (i, j)::rem_sdup ((i', j')::xs) + fun sexec [] a = a + | sexec ((i, j)::ops) a = sexec ops (swap i j a) + + fun swaps_of (l:int list) = + ops_of 0 (List.length l - 1) (fromList l) |> rem_sdup + |> filter (fn (i, j) => i <> j) +end + +*} + +text {* Examples *} + +ML {* + val l = [8, 9, 10, 1, 12, 13, 14] + val ops = (swaps_of l) + val a = (sexec ops (Array.fromList l)) + val l' = Array.vector a + val a = sexec (rev ops) a +*} + +end \ No newline at end of file diff -r 77daf1b85cf0 -r a5f5b9336007 thys2/Sort_ops.thy~ --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/Sort_ops.thy~ Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,65 @@ +theory Sort_ops +imports Main +begin + +ML {* +local + open Array +in + + fun swap i j a = let val ai = sub(a, i); + val _ = update(a, i, sub(a, j)) + val _ = update(a, j, ai) + in + a + end + + fun pre_ops_of a = + getM :|-- (fn (l, r, piv, i, j) => let + (* val _ = Output.tracing ("("^string_of_int i^","^string_of_int j^")") *) + in + if (j < i) then (returnM (swap l j a) |-- returnM [(l, j)]) + else (if (sub(a, i) <= piv andalso i <= r) then (setM (l, r, piv, i + 1, j) + |-- pre_ops_of a) + else if (sub(a, j) > piv ) then (setM (l, r, piv, i, j - 1) |-- pre_ops_of a) + else (pre_ops_of (swap i j a) :|-- (fn ops => returnM ((i,j)::ops)))) + end + ) + + fun ops_of i j a = + if (j < i) then [] + else let + val piv = sub(a, i) + val (ops1, (_, _, _, i', j')) = pre_ops_of a (i, j, piv, i, j) |> normVal + val ops2 = ops_of i (j' - 1) a + val ops3 = ops_of (j' + 1) j a + in + ops1 @ ops2 @ ops3 + end + + fun rem_sdup [] = [] + | rem_sdup [c] = [c] + | rem_sdup ((i, j)::(i', j')::xs) = if ((i = i' andalso j = j') orelse + (i = j' andalso j = i')) then rem_sdup (xs) + else (i, j)::rem_sdup ((i', j')::xs) + fun sexec [] a = a + | sexec ((i, j)::ops) a = sexec ops (swap i j a) + + fun swaps_of (l:int list) = + ops_of 0 (List.length l - 1) (fromList l) |> rem_sdup + |> filter (fn (i, j) => i <> j) +end + +*} + +text {* Examples *} + +ML {* + val l = [8, 9, 10, 1, 12, 13, 14] + val ops = (swaps_of l) + val a = (sexec ops (Array.fromList l)) + val l' = Array.vector a + val a = sexec (rev ops) a +*} + +end \ No newline at end of file diff -r 77daf1b85cf0 -r a5f5b9336007 thys2/Subgoal.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/Subgoal.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,76 @@ +theory Subgoal +imports Term_pat +begin + +section {* Monad for goal and subgoal manipulation, useful for debugging tactics *} + +subsection {* Implementation *} + +ML {* + fun goalM raw_goal = m0M' (fn ((i, _), ctxt) => + let + val thy = Proof_Context.theory_of ctxt + val goal = raw_goal |> cterm_of thy |> Goal.init + (* val ctxt' = Variable.auto_fixes raw_goal ctxt *) + val ctxt' = ctxt + in (s2M ctxt' |-- (s1M (i, goal)) |-- returnM goal) end) + +local + fun pr_goal ctxt goal = Goal_Display.pretty_goals ctxt goal |> + Pretty.markup_chunks Markup.state |> Pretty.writeln +in + fun tacM tac = m1M' (fn (i, goal) => m2M (fn ctxt => + let + val goal' = (Seq.hd (tac ctxt goal)) + val _ = pr_goal ctxt goal' + in + s1M (i, goal') |-- returnM goal' + end)) +end + + fun focusM i = pushM |-- (liftM (m0M (fn ((_, goal), ctxt) => + let val (focus as {context = ctxt', ...}, goal') = Subgoal.focus ctxt i goal + in + setM ((i, goal'), ctxt') |-- returnM focus + end))) + + fun end_focusM x = ( popM + :|-- (fn ((i, goal'), _) => topM (fn ((old_i, goal), ctxt) => + let val ({context = ctxt', params, asms, ...}, _) = + Subgoal.focus ctxt i goal + val new_goal = Seq.lifts (Subgoal.retrofit ctxt' ctxt params asms i) + (Seq.single goal') goal |> Seq.hd + in s1M (old_i, new_goal) |-- returnM new_goal end ))) x +*} + +subsection {* Examples *} + +lemma init: "A \ A" + by simp + +ML {* + val r as (result, _)= [((0, @{thm "init"}), @{context})] |> + @{fterm "(A \ B \ A)"} + :|-- goalM + |-- focusM 1 + :|--(fn {prems, ...} => let + val conj = hd prems + val r = rtac (conj RS @{thm HOL.conjunct1}) 1 + in tacM (K r) end ) + |-- end_focusM >> Drule.export_without_context |> normVal +*} + + +ML {* + val r = [((0, @{thm "init"}), @{context})] |> + @{fterm "(A \ B \ A)"} + :|-- goalM + |-- tacM (fn ctxt => Subgoal.FOCUS (fn {prems, ...} => let + val conj = hd prems + val r = rtac (conj RS @{thm HOL.conjunct1}) 1 + in r end ) ctxt 1) >> Drule.export_without_context + |> normVal +*} + +end + diff -r 77daf1b85cf0 -r a5f5b9336007 thys2/Term_pat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/Term_pat.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,244 @@ +theory Term_pat +imports Main +begin + +text {* + @{text "trace_match"} is the flag to control the tracing of matching activities. +*} + +ML {* + val trace_match = Attrib.setup_config_bool @{binding trace_match} (K false) +*} + +section {* Monad with support of excepion *} + +ML {* + +structure Exceptional = +struct + datatype ('a, 'b) Exceptional = + Exception of 'a | Success of 'b + fun (scan1 :|-- scan2) st = + case scan1 st of + Success (v, st') => scan2 v st' + | Exception e => Exception e + fun returnM v = (fn st => Success (v, st)) + fun (scan1 |-- scan2) = scan1 :|-- (K scan2) + fun (scan1 -- scan2) = scan1 :|-- (fn v1 => scan2 :|-- (fn v2 => returnM (v1, v2))) + fun liftE scan st = let + val (v, st') = scan st + in Success (v, st') end + fun getM y = Success (y, y) + fun setM v = (fn _ => Success (v, v)) + fun failM x = (fn _ => raise (Scan.fail x)) + fun m0M f = getM :|-- f + fun pushM x = (m0M (fn (top::rest) => setM (top::top::rest) |-- returnM top)) x + fun popM x = (m0M (fn (top::rest) => setM rest |-- returnM top)) x + (* [f] is assumed to be a monad valued function, so that it can be constructed using monad combinators. + For this reason, [f] is assumed to take two arguments. + *) + fun topM f = m0M (fn (top::rest) => + case f top top of + Success (v, top') => (setM (top'::rest) |-- returnM v) + | Exception e => (K (Exception e))) + (* all the [f] in the following are monad or moand function on stack element (not on stack itself )*) + (* [f] in the following is assumed to be plain monad on stack element, not a monad valued function *) + fun liftM f = topM (K f) + (* [f] in the following is a monad valued funtion depending only on the first component *) + fun m1M f = m0M (fn (st1, st2) => f st1) + (* [f] in the following is a monad valued funtion depending only on the second component *) + fun m2M f = m0M (fn (st1, st2) => f st2) + fun localM f = pushM |-- f :|-- (fn res => popM |-- returnM res) + fun s1M n1 = m0M (fn (_, st2) => setM (n1, st2)) + fun s2M n2 = m0M (fn (st1, _) => setM (st1, n2)) + (* The flowing primed version is meant to substitute the unprimed version *) + fun m0M' f = liftM (m0M f) + fun m1M' f = liftM (m1M f) + fun m2M' f = liftM (m2M f) + fun lift scan = (fn (d, a) => + case scan a of + Success (v, b) => Success (v, (d, b)) + | Exception e => Exception e + ) + fun normVal (Success v) = v + fun normExcp (Exception e) = e + fun (scan >> f) = scan :|-- (fn v => returnM (f v)) + fun raiseM v = K (Exception v) +end +*} + +(* ccc *) + +section {* State monad with stack *} + +ML {* +structure StackMonad = struct + (* getM fetches the underlying state *) + fun getM y = (y, y) + (* setM sets the underlying state *) + fun setM v = (fn _ => (v, v)) + (* returnM retruns any value *) + fun returnM v = (fn st => (v, st)) + (* failM failes the current process *) + fun failM x = (fn _ => raise (Scan.fail x)) + (* [m0M f] applies function f to the underlying state *) + fun m0M f = getM :|-- f + fun pushM x = (m0M (fn (top::rest) => setM (top::top::rest) |-- returnM top)) x + fun popM x = (m0M (fn (top::rest) => setM rest |-- returnM top)) x + (* [f] is assumed to be a monad valued function, so that it can be constructed using monad combinators. + For this reason, [f] is assumed to take two arguments. + *) + fun topM f = m0M (fn (top::rest) => let val (v, top') = f top top in + setM (top'::rest) |-- returnM v + end) + (* all the [f] in the following are monad or moand function on stack element (not on stack itself )*) + (* [f] in the following is assumed to be plain monad on stack element, not a monad valued function *) + fun liftM f = topM (K f) + (* [f] in the following is a monad valued funtion depending only on the first component *) + fun m1M f = m0M (fn (st1, st2) => f st1) + (* [f] in the following is a monad valued funtion depending only on the second component *) + fun m2M f = m0M (fn (st1, st2) => f st2) + fun localM f = pushM |-- f :|-- (fn res => popM |-- returnM res) + fun s1M n1 = m0M (fn (_, st2) => setM (n1, st2)) + fun s2M n2 = m0M (fn (st1, _) => setM (st1, n2)) + fun dgM scan arg = let val (v, [arg']) = scan [arg] in (v, arg') end + (* The flowing primed version is meant to substitute the unprimed version *) + fun m0M' f = liftM (m0M f) + fun m1M' f = liftM (m1M f) + fun m2M' f = liftM (m2M f) + val liftE = I + val lift = Scan.lift + val :|-- = Scan.:|-- + val |-- = Scan.|-- + val -- = Scan.-- + val >> = Scan.>> + val normVal = I +end +*} + +ML {* + open StackMonad + open Exceptional +*} + +section {* State monad with stack *} + +ML {* +local + val tracing = (fn ctxt => fn str => + if (Config.get ctxt trace_match) then tracing str else ()) + fun string_of_term ctxt t = t |> Syntax.pretty_term ctxt |> Pretty.str_of + fun pterm ctxt t = + t |> string_of_term ctxt |> tracing ctxt +in + fun tracingM s = liftM (m2M (fn ctxt => returnM (tracing ctxt s))) + fun ptermM t = liftM (m2M (fn ctxt => returnM (pterm ctxt t))) + fun presultM tag scan = scan :|-- (fn res => + tracingM tag |-- + ptermM res |-- + returnM res + ) +end +*} + +ML {* +local + val tracing = (fn ctxt => fn str => + if (Config.get ctxt trace_match) then tracing str else ()) + fun string_of_term ctxt t = t |> Syntax.pretty_term ctxt |> Pretty.str_of + fun pterm ctxt t = + t |> string_of_term ctxt |> tracing ctxt + fun match_fun1 pat term (st, ctxt) = let + val _ = tracing ctxt "match_fun on:" + val _ = tracing ctxt pat + val _ = pterm ctxt term + in + (Proof_Context.match_bind_i true + [([Proof_Context.read_term_pattern ctxt pat], term)] ctxt |> + (fn (_, ctxt) => (term, (st, ctxt)))) + handle (ERROR x) => + raise (Scan.fail x) + end +in + fun match_fun pat term = liftE (match_fun1 pat term) +end +*} + +ML {* + fun err_trans (ERROR x) = raise (Scan.fail x) + | err_trans e = e +*} + +ML {* + fun exnM f monad = (fn inp => monad inp handle e => raise (f e)) +*} + +ML {* + fun match_fun pat term = + tracingM "match_fun on:" |-- + tracingM pat |-- + ptermM term |-- + liftM (lift (m0M (fn ctxt => (* setM (Variable.auto_fixes term ctxt) |-- *) + (liftE (Proof_Context.match_bind_i true [([Proof_Context.read_term_pattern ctxt pat], term)])) |-- + returnM term + )) |> exnM err_trans) +*} + +ML {* +local + val >> = Scan.>> +in + val match_parser = + (Args.name_source) >> (fn s => "(match_fun "^ML_Syntax.print_string s^")") +end +*} + +ML {* +local + val >> = Scan.>> +in + val fterm_parser = + Args.name_source >> + (fn term => "liftM (exnM err_trans (lift (m0M (fn ctxt => returnM ((Syntax.read_term ctxt) "^ML_Syntax.print_string term ^")))))") +end +*} + +setup {* + ML_Antiquote.inline @{binding match} (Scan.lift match_parser) #> + ML_Antiquote.inline @{binding fterm} (Scan.lift fterm_parser) +*} + +ML {* + infixr 0 ||| + fun (scan1 ||| scan2) xs = scan1 xs handle _ => scan2 xs +*} + +declare [[trace_match = false]] + +ML {* + val t = [((), @{context})] |> + (@{match "(?x::nat) + ?y"} @{term "(1::nat) + 2"} |-- + @{fterm "?x"}) +*} + + +ML {* (* testings and examples *) + fun trans t = (@{match "(?x ::nat) + ?y"} t |-- + (@{fterm "?x"} :|-- trans) :|-- @{match "?u"} |-- + (@{fterm "?y"} :|-- trans) :|-- @{match "?v"} |-- + @{fterm "?v * ?v * ?u"} + || + @{match "(?x ::nat) - ?y"} t |-- + (@{fterm "?x"} :|-- trans) :|-- @{match "?u"} |-- + (@{fterm "?y"} :|-- trans) :|-- @{match "v"} |-- + @{fterm "?v - ?v * ?u"} + || + (returnM t)) +*} + +ML {* + val t = trans @{term "(u - (x::nat) + y)"} [((), @{context})] |> normVal |> + apfst (cterm_of @{theory}) |> fst + *} + +end \ No newline at end of file diff -r 77daf1b85cf0 -r a5f5b9336007 thys2/Thm_inst.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/Thm_inst.thy Sat Sep 13 10:07:14 2014 +0800 @@ -0,0 +1,130 @@ +theory Thm_inst +imports Term_pat +begin + +section {* Monad for theorem instantiation based on bindings in proof contexts *} + +subsection {* Implementation *} + +ML {* +local +fun read_termTs ctxt schematic ts Ts = + let + val ts' = + map2 (Type.constraint o Type_Infer.paramify_vars) Ts ts + |> Syntax.check_terms ((schematic ? Proof_Context.set_mode Proof_Context.mode_schematic) ctxt) + |> Variable.polymorphic ctxt; + val Ts' = map Term.fastype_of ts'; + val tyenv = fold Type.raw_match (Ts ~~ Ts') Vartab.empty; + in (ts', map (apsnd snd) (Vartab.dest tyenv)) end; + +fun make_inst f v = + let + val t = Var v; + val t' = f t; + in if t aconv t' then NONE else SOME (t, t') end; + + +fun error_var msg xi = error (msg ^ quote (Term.string_of_vname xi)); + +fun the_type vars (xi: indexname) = + (case AList.lookup (op =) vars xi of + SOME T => T + | NONE => error_var "No such variable in theorem: " xi); + +fun instantiate inst = + Term_Subst.instantiate ([], map (fn (xi, t) => ((xi, Term.fastype_of t), t)) inst) #> + Envir.beta_norm; + +fun make_instT f v = + let + val T = TVar v; + val T' = f T; + in if T = T' then NONE else SOME (T, T') end; + +fun read_insts ctxt term_insts (tvars, vars) = + let + val thy = Proof_Context.theory_of ctxt; + val cert = Thm.cterm_of thy; + val certT = Thm.ctyp_of thy; + + + (* type instantiations *) + + val instT1 = Term_Subst.instantiateT []; + val vars1 = map (apsnd instT1) vars; + + + (* term instantiations *) + + val (xs, ts) = split_list term_insts; + val Ts = map (the_type vars1) xs; + val (ts, inferred) = read_termTs ctxt false ts Ts; + + val instT2 = Term.typ_subst_TVars inferred; + val vars2 = map (apsnd instT2) vars1; + val inst2 = instantiate (xs ~~ ts); + + + (* result *) + + val inst_tvars = map_filter (make_instT (instT2 o instT1)) tvars; + val inst_vars = map_filter (make_inst inst2) vars2; + in + (map (pairself certT) inst_tvars, map (pairself cert) inst_vars) + end + + val add_used = + (Thm.fold_terms o fold_types o fold_atyps) + (fn TFree (a, _) => insert (op =) a + | TVar ((a, _), _) => insert (op =) a + | _ => I) + +fun term_instantiate ctxt term_insts thm = + let + val ctxt' = ctxt + |> Variable.declare_thm thm + |> fold (fn a => Variable.declare_names (Logic.mk_type (TFree (a, dummyS)))) (add_used thm []); + (* FIXME !? *) + val tvars = Thm.fold_terms Term.add_tvars thm []; + val vars = Thm.fold_terms Term.add_vars thm []; + val insts = read_insts ctxt' term_insts (tvars, vars); + in + Drule.instantiate_normalize insts thm + |> Rule_Cases.save thm + end + +in + fun ctxt_instantiate ctxt thm = let + val binds = Variable.binds_of ctxt + val vars = Thm.fold_terms Term.add_vars thm [] |> map fst + val inst = map_filter ( Vartab.lookup_key binds) vars |> map (apsnd snd) + val thm' = term_instantiate ctxt inst thm + in thm' end + + fun thm_instM thm = + liftM (lift (m0M (fn ctxt => + returnM (ctxt_instantiate ctxt thm) + )) |> exnM err_trans) + +end +*} + +subsection {* Examples *} + +lemma h: "x = x \ y = y" + by simp + +ML {* + val ctxt = @{context} + val t = ctxt_instantiate ctxt @{thm h} |> prop_of +*} + +ML {* + val t = [((), @{context})] |> @{match "?x"} @{term "[]::nat list"} |-- + @{match "?y"} @{term "1::nat"} |-- + thm_instM @{thm h} +*} + +end +