diff --git a/thys/MFOTL_Checker/Checker.thy b/thys/MFOTL_Checker/Checker.thy new file mode 100644 --- /dev/null +++ b/thys/MFOTL_Checker/Checker.thy @@ -0,0 +1,3320 @@ +(*<*) +theory Checker + imports Prelim Proof_System Proof_Object "HOL-Library.Simps_Case_Conv" +begin +(*>*) + +section \Proof Checker\ + +unbundle MFOTL_notation + +context fixes \ :: "('n, 'd :: {default, linorder}) trace" + +begin + +fun s_check :: "('n, 'd) env \ ('n, 'd) formula \ ('n, 'd) sproof \ bool" +and v_check :: "('n, 'd) env \ ('n, 'd) formula \ ('n, 'd) vproof \ bool" where + "s_check v f p = (case (f, p) of + (\, STT i) \ True + | (r \ ts, SPred i s ts') \ + (r = s \ ts = ts' \ (r, v\<^bold>\ts\<^bold>\) \ \ \ i) + | (x \<^bold>\ c, SEq_Const i x' c') \ + (c = c' \ x = x' \ v x = c) + | (\\<^sub>F \, SNeg vp) \ v_check v \ vp + | (\ \\<^sub>F \, SOrL sp1) \ s_check v \ sp1 + | (\ \\<^sub>F \, SOrR sp2) \ s_check v \ sp2 + | (\ \\<^sub>F \, SAnd sp1 sp2) \ s_check v \ sp1 \ s_check v \ sp2 \ s_at sp1 = s_at sp2 + | (\ \\<^sub>F \, SImpL vp1) \ v_check v \ vp1 + | (\ \\<^sub>F \, SImpR sp2) \ s_check v \ sp2 + | (\ \\<^sub>F \, SIffSS sp1 sp2) \ s_check v \ sp1 \ s_check v \ sp2 \ s_at sp1 = s_at sp2 + | (\ \\<^sub>F \, SIffVV vp1 vp2) \ v_check v \ vp1 \ v_check v \ vp2 \ v_at vp1 = v_at vp2 + | (\\<^sub>Fx. \, SExists y val sp) \ (x = y \ s_check (v (x := val)) \ sp) + | (\\<^sub>Fx. \, SForall y sp_part) \ (let i = s_at (part_hd sp_part) + in x = y \ (\(sub, sp) \ SubsVals sp_part. s_at sp = i \ (\z \ sub. s_check (v (x := z)) \ sp))) + | (\<^bold>Y I \, SPrev sp) \ + (let j = s_at sp; i = s_at (SPrev sp) in + i = j+1 \ mem (\ \ i) I \ s_check v \ sp) + | (\<^bold>X I \, SNext sp) \ + (let j = s_at sp; i = s_at (SNext sp) in + j = i+1 \ mem (\ \ j) I \ s_check v \ sp) + | (\<^bold>P I \, SOnce i sp) \ + (let j = s_at sp in + j \ i \ mem (\ \ i - \ \ j) I \ s_check v \ sp) + | (\<^bold>F I \, SEventually i sp) \ + (let j = s_at sp in + j \ i \ mem (\ \ j - \ \ i) I \ s_check v \ sp) + | (\<^bold>H I \, SHistoricallyOut i) \ + \ \ i < \ \ 0 + left I + | (\<^bold>H I \, SHistorically i li sps) \ + (li = (case right I of \ \ 0 | enat b \ ETP \ (\ \ i - b)) + \ \ \ 0 + left I \ \ \ i + \ map s_at sps = [li ..< (LTP_p \ i I) + 1] + \ (\sp \ set sps. s_check v \ sp)) + | (\<^bold>G I \, SAlways i hi sps) \ + (hi = (case right I of enat b \ LTP_f \ i b) + \ right I \ \ + \ map s_at sps = [(ETP_f \ i I) ..< hi + 1] + \ (\sp \ set sps. s_check v \ sp)) + | (\ \<^bold>S I \, SSince sp2 sp1s) \ + (let i = s_at (SSince sp2 sp1s); j = s_at sp2 in + j \ i \ mem (\ \ i - \ \ j) I + \ map s_at sp1s = [j+1 ..< i+1] + \ s_check v \ sp2 + \ (\sp1 \ set sp1s. s_check v \ sp1)) + | (\ \<^bold>U I \, SUntil sp1s sp2) \ + (let i = s_at (SUntil sp1s sp2); j = s_at sp2 in + j \ i \ mem (\ \ j - \ \ i) I + \ map s_at sp1s = [i ..< j] \ s_check v \ sp2 + \ (\sp1 \ set sp1s. s_check v \ sp1)) + | ( _ , _) \ False)" +| "v_check v f p = (case (f, p) of + (\, VFF i) \ True + | (r \ ts, VPred i pred ts') \ + (r = pred \ ts = ts' \ (r, v\<^bold>\ts\<^bold>\) \ \ \ i) + | (x \<^bold>\ c, VEq_Const i x' c') \ + (c = c' \ x = x' \ v x \ c) + | (\\<^sub>F \, VNeg sp) \ s_check v \ sp + | (\ \\<^sub>F \, VOr vp1 vp2) \ v_check v \ vp1 \ v_check v \ vp2 \ v_at vp1 = v_at vp2 + | (\ \\<^sub>F \, VAndL vp1) \ v_check v \ vp1 + | (\ \\<^sub>F \, VAndR vp2) \ v_check v \ vp2 + | (\ \\<^sub>F \, VImp sp1 vp2) \ s_check v \ sp1 \ v_check v \ vp2 \ s_at sp1 = v_at vp2 + | (\ \\<^sub>F \, VIffSV sp1 vp2) \ s_check v \ sp1 \ v_check v \ vp2 \ s_at sp1 = v_at vp2 + | (\ \\<^sub>F \, VIffVS vp1 sp2) \ v_check v \ vp1 \ s_check v \ sp2 \ v_at vp1 = s_at sp2 + | (\\<^sub>Fx. \, VExists y vp_part) \ (let i = v_at (part_hd vp_part) + in x = y \ (\(sub, vp) \ SubsVals vp_part. v_at vp = i \ (\z \ sub. v_check (v (x := z)) \ vp))) + | (\\<^sub>Fx. \, VForall y val vp) \ (x = y \ v_check (v (x := val)) \ vp) + | (\<^bold>Y I \, VPrev vp) \ + (let j = v_at vp; i = v_at (VPrev vp) in + i = j+1 \ v_check v \ vp) + | (\<^bold>Y I \, VPrevZ) \ True + | (\<^bold>Y I \, VPrevOutL i) \ + i > 0 \ \ \ i < left I + | (\<^bold>Y I \, VPrevOutR i) \ + i > 0 \ enat (\ \ i) > right I + | (\<^bold>X I \, VNext vp) \ + (let j = v_at vp; i = v_at (VNext vp) in + j = i+1 \ v_check v \ vp) + | (\<^bold>X I \, VNextOutL i) \ + \ \ (i+1) < left I + | (\<^bold>X I \, VNextOutR i) \ + enat (\ \ (i+1)) > right I + | (\<^bold>P I \, VOnceOut i) \ + \ \ i < \ \ 0 + left I + | (\<^bold>P I \, VOnce i li vps) \ + (li = (case right I of \ \ 0 | enat b \ ETP_p \ i b) + \ \ \ 0 + left I \ \ \ i + \ map v_at vps = [li ..< (LTP_p \ i I) + 1] + \ (\vp \ set vps. v_check v \ vp)) + | (\<^bold>F I \, VEventually i hi vps) \ + (hi = (case right I of enat b \ LTP_f \ i b) \ right I \ \ + \ map v_at vps = [(ETP_f \ i I) ..< hi + 1] + \ (\vp \ set vps. v_check v \ vp)) + | (\<^bold>H I \, VHistorically i vp) \ + (let j = v_at vp in + j \ i \ mem (\ \ i - \ \ j) I \ v_check v \ vp) + | (\<^bold>G I \, VAlways i vp) \ + (let j = v_at vp + in j \ i \ mem (\ \ j - \ \ i) I \ v_check v \ vp) + | (\ \<^bold>S I \, VSinceOut i) \ + \ \ i < \ \ 0 + left I + | (\ \<^bold>S I \, VSince i vp1 vp2s) \ + (let j = v_at vp1 in + (case right I of \ \ True | enat b \ ETP_p \ i b \ j) \ j \ i + \ \ \ 0 + left I \ \ \ i + \ map v_at vp2s = [j ..< (LTP_p \ i I) + 1] \ v_check v \ vp1 + \ (\vp2 \ set vp2s. v_check v \ vp2)) + | (\ \<^bold>S I \, VSinceInf i li vp2s) \ + (li = (case right I of \ \ 0 | enat b \ ETP_p \ i b) + \ \ \ 0 + left I \ \ \ i + \ map v_at vp2s = [li ..< (LTP_p \ i I) + 1] + \ (\vp2 \ set vp2s. v_check v \ vp2)) + | (\ \<^bold>U I \, VUntil i vp2s vp1) \ + (let j = v_at vp1 in + (case right I of \ \ True | enat b \ j < LTP_f \ i b) \ i \ j + \ map v_at vp2s = [ETP_f \ i I ..< j + 1] \ v_check v \ vp1 + \ (\vp2 \ set vp2s. v_check v \ vp2)) + | (\ \<^bold>U I \, VUntilInf i hi vp2s) \ + (hi = (case right I of enat b \ LTP_f \ i b) \ right I \ \ + \ map v_at vp2s = [ETP_f \ i I ..< hi + 1] + \ (\vp2 \ set vp2s. v_check v \ vp2)) + | ( _ , _) \ False)" + +declare s_check.simps[simp del] v_check.simps[simp del] +simps_of_case s_check_simps[simp]: s_check.simps[unfolded prod.case] (splits: formula.split sproof.split) +simps_of_case v_check_simps[simp]: v_check.simps[unfolded prod.case] (splits: formula.split vproof.split) + +subsection \Checker Soundness\ + +lemma check_soundness: + "s_check v \ sp \ SAT \ v (s_at sp) \" + "v_check v \ vp \ VIO \ v (v_at vp) \" +proof (induction sp and vp arbitrary: v \ and v \) + case STT + then show ?case by (cases \) (auto intro: SAT_VIO.STT) +next + case SPred + then show ?case by (cases \) (auto intro: SAT_VIO.SPred) +next + case SEq_Const + then show ?case by (cases \) (auto intro: SAT_VIO.SEq_Const) +next + case SNeg + then show ?case by (cases \) (auto intro: SAT_VIO.SNeg) +next + case SAnd + then show ?case by (cases \) (auto intro: SAT_VIO.SAnd) +next + case SOrL + then show ?case by (cases \) (auto intro: SAT_VIO.SOrL) +next + case SOrR + then show ?case by (cases \) (auto intro: SAT_VIO.SOrR) +next + case SImpR + then show ?case by (cases \) (auto intro: SAT_VIO.SImpR) +next + case SImpL + then show ?case by (cases \) (auto intro: SAT_VIO.SImpL) +next + case SIffSS + then show ?case by (cases \) (auto intro: SAT_VIO.SIffSS) +next + case SIffVV + then show ?case by (cases \) (auto intro: SAT_VIO.SIffVV) +next + case (SExists x z sp) + with SExists(1)[of "v(x := z)"] show ?case + by (cases \) (auto intro: SAT_VIO.SExists) +next + case (SForall x part) + then show ?case + proof (cases \) + case (Forall y \) + show ?thesis unfolding Forall + proof (intro SAT_VIO.SForall allI) + fix z + let ?sp = "lookup_part part z" + from lookup_part_SubsVals[of z part] obtain D where "z \ D" "(D, ?sp) \ SubsVals part" + by blast + with SForall(2-) Forall have "s_check (v(y := z)) \ ?sp" "s_at ?sp = s_at (SForall x part)" + by auto + then show "SAT \ (v(y := z)) (s_at (SForall x part)) \" + by (auto simp del: fun_upd_apply dest!: SForall(1)[rotated]) + qed + qed auto +next + case (SSince spsi sps) + then show ?case + proof (cases \) + case (Since \ I \) + show ?thesis + using SSince(3) + unfolding Since + proof (intro SAT_VIO.SSince[of "s_at spsi"], goal_cases le mem SAT\ SAT\) + case (SAT\ k) + then show ?case + by (cases "k \ s_at (hd sps)") + (auto 0 3 simp: Let_def elim: map_setE[of _ _ _ k] intro: SSince(2) dest!: sym[of "s_at _" "Suc (s_at _)"]) + qed (auto simp: Let_def intro: SSince(1)) + qed auto +next + case (SOnce i sp) + then show ?case + proof (cases \) + case (Once I \) + show ?thesis + using SOnce + unfolding Once + by (intro SAT_VIO.SOnce[of "s_at sp"]) (auto simp: Let_def) + qed auto +next + case (SEventually i sp) + then show ?case + proof (cases \) + case (Eventually I \) + show ?thesis + using SEventually + unfolding Eventually + by (intro SAT_VIO.SEventually[of _ "s_at sp"]) (auto simp: Let_def) + qed auto +next + case SHistoricallyOut + then show ?case by (cases \) (auto intro: SAT_VIO.SHistoricallyOut) +next + case (SHistorically i li sps) + then show ?case + proof (cases \) + case (Historically I \) + {fix k + define j where j_def: "j \ case right I of \ \ 0 | enat n \ ETP \ (\ \ i - n)" + assume k_def: "k \ j \ k \ i \ k \ LTP \ (\ \ i - left I)" + from SHistorically Historically j_def have map: "set (map s_at sps) = set [j ..< Suc (LTP_p \ i I)]" + by (auto simp: Let_def) + then have kset: "k \ set ([j ..< Suc (LTP_p \ i I)])" using j_def k_def by auto + then obtain x where x: "x \ set sps" "s_at x = k" using k_def map + unfolding set_map set_eq_iff image_iff + by metis + then have "SAT \ v k \" using SHistorically unfolding Historically + by (auto simp: Let_def) + } note * = this + show ?thesis + using SHistorically * + unfolding Historically + by (auto simp: Let_def intro!: SAT_VIO.SHistorically) + qed (auto intro: SAT_VIO.intros) +next + case (SAlways i hi sps) + then show ?case + proof (cases \) + case (Always I \) + obtain n where n_def: "right I = enat n" + using SAlways + by (auto simp: Always split: enat.splits) + {fix k + define j where j_def: "j \ LTP \ (\ \ i + n)" + assume k_def: "k \ j \ k \ i \ k \ ETP \ (\ \ i + left I)" + from SAlways Always j_def have map: "set (map s_at sps) = set [(ETP_f \ i I) ..< Suc j]" + by (auto simp: Let_def n_def) + then have kset: "k \ set ([(ETP_f \ i I) ..< Suc j])" using k_def j_def by auto + then obtain x where x: "x \ set sps" "s_at x = k" using k_def map + unfolding set_map set_eq_iff image_iff + by metis + then have "SAT \ v k \" using SAlways unfolding Always + by (auto simp: Let_def n_def) + } note * = this + then show ?thesis + using SAlways + unfolding Always + by (auto simp: Let_def n_def intro: SAT_VIO.SAlways split: if_splits enat.splits) + qed(auto intro: SAT_VIO.intros) +next + case (SUntil sps spsi) + then show ?case + proof (cases \) + case (Until \ I \) + show ?thesis + using SUntil(3) + unfolding Until + proof (intro SAT_VIO.SUntil[of _ "s_at spsi"], goal_cases le mem SAT\ SAT\) + case (SAT\ k) + then show ?case + by (cases "k \ s_at (hd sps)") + (auto 0 3 simp: Let_def elim: map_setE[of _ _ _ k] intro: SUntil(1)) + qed (auto simp: Let_def intro: SUntil(2)) + qed auto +next + case (SNext sp) + then show ?case by (cases \) (auto simp add: Let_def SAT_VIO.SNext) +next + case (SPrev sp) + then show ?case by (cases \) (auto simp add: Let_def SAT_VIO.SPrev) +next + case VFF + then show ?case by (cases \) (auto intro: SAT_VIO.VFF) +next + case VPred + then show ?case by (cases \) (auto intro: SAT_VIO.VPred) +next + case VEq_Const + then show ?case by (cases \) (auto intro: SAT_VIO.VEq_Const) +next + case VNeg + then show ?case by (cases \) (auto intro: SAT_VIO.VNeg) +next + case VOr + then show ?case by (cases \) (auto intro: SAT_VIO.VOr) +next + case VAndL + then show ?case by (cases \) (auto intro: SAT_VIO.VAndL) +next + case VAndR + then show ?case by (cases \) (auto intro: SAT_VIO.VAndR) +next + case VImp + then show ?case by (cases \) (auto intro: SAT_VIO.VImp) +next + case VIffSV + then show ?case by (cases \) (auto intro: SAT_VIO.VIffSV) +next + case VIffVS + then show ?case by (cases \) (auto intro: SAT_VIO.VIffVS) +next + case (VExists x part) + then show ?case + proof (cases \) + case (Exists y \) + show ?thesis unfolding Exists + proof (intro SAT_VIO.VExists allI) + fix z + let ?vp = "lookup_part part z" + from lookup_part_SubsVals[of z part] obtain D where "z \ D" "(D, ?vp) \ SubsVals part" + by blast + with VExists(2-) Exists have "v_check (v(y := z)) \ ?vp" "v_at ?vp = v_at (VExists x part)" + by auto + then show "VIO \ (v(y := z)) (v_at (VExists x part)) \" + by (auto simp del: fun_upd_apply dest!: VExists(1)[rotated]) + qed + qed auto +next + case (VForall x z vp) + with VForall(1)[of "v(x := z)"] show ?case + by (cases \) (auto intro: SAT_VIO.VForall) +next + case VOnceOut + then show ?case by (cases \) (auto intro: SAT_VIO.VOnceOut) +next + case (VOnce i li vps) + then show ?case + proof (cases \) + case (Once I \) + {fix k + define j where j_def: "j \ case right I of \ \ 0 | enat n \ ETP \ (\ \ i - n)" + assume k_def: "k \ j \ k \ i \ k \ LTP \ (\ \ i - left I)" + from VOnce Once j_def have map: "set (map v_at vps) = set [j ..< Suc (LTP_p \ i I)]" + by (auto simp: Let_def) + then have kset: "k \ set ([j ..< Suc (LTP_p \ i I)])" using j_def k_def by auto + then obtain x where x: "x \ set vps" "v_at x = k" using k_def map + unfolding set_map set_eq_iff image_iff + by metis + then have "VIO \ v k \" using VOnce unfolding Once + by (auto simp: Let_def) + } note * = this + show ?thesis + using VOnce * + unfolding Once + by (auto simp: Let_def intro!: SAT_VIO.VOnce) + qed (auto intro: SAT_VIO.intros) +next + case (VEventually i hi vps) + then show ?case + proof (cases \) + case (Eventually I \) + obtain n where n_def: "right I = enat n" + using VEventually + by (auto simp: Eventually split: enat.splits) + {fix k + define j where j_def: "j \ LTP \ (\ \ i + n)" + assume k_def: "k \ j \ k \ i \ k \ ETP \ (\ \ i + left I)" + from VEventually Eventually j_def have map: "set (map v_at vps) = set [(ETP_f \ i I) ..< Suc j]" + by (auto simp: Let_def n_def) + then have kset: "k \ set ([(ETP_f \ i I) ..< Suc j])" using k_def j_def by auto + then obtain x where x: "x \ set vps" "v_at x = k" using k_def map + unfolding set_map set_eq_iff image_iff + by metis + then have "VIO \ v k \" using VEventually unfolding Eventually + by (auto simp: Let_def n_def) + } note * = this + then show ?thesis + using VEventually + unfolding Eventually + by (auto simp: Let_def n_def intro: SAT_VIO.VEventually split: if_splits enat.splits) + qed(auto intro: SAT_VIO.intros) +next + case (VHistorically i vp) + then show ?case + proof (cases \) + case (Historically I \) + show ?thesis + using VHistorically + unfolding Historically + by (intro SAT_VIO.VHistorically[of "v_at vp"]) (auto simp: Let_def) + qed auto +next + case (VAlways i vp) + then show ?case + proof (cases \) + case (Always I \) + show ?thesis + using VAlways + unfolding Always + by (intro SAT_VIO.VAlways[of _ "v_at vp"]) (auto simp: Let_def) + qed auto +next + case VNext + then show ?case by (cases \) (auto intro: SAT_VIO.VNext) +next + case VNextOutR + then show ?case by (cases \) (auto intro: SAT_VIO.VNextOutR) +next + case VNextOutL + then show ?case by (cases \) (auto intro: SAT_VIO.VNextOutL) +next + case VPrev + then show ?case by (cases \) (auto intro: SAT_VIO.VPrev) +next + case VPrevOutR + then show ?case by (cases \) (auto intro: SAT_VIO.VPrevOutR) +next + case VPrevOutL + then show ?case by (cases \) (auto intro: SAT_VIO.VPrevOutL) +next + case VPrevZ + then show ?case by (cases \) (auto intro: SAT_VIO.VPrevZ) +next + case VSinceOut + then show ?case by (cases \) (auto intro: SAT_VIO.VSinceOut) +next + case (VSince i vp vps) + then show ?case + proof (cases \) + case (Since \ I \) + {fix k + assume k_def: "k \ v_at vp \ k \ i \ k \ LTP \ (\ \ i - left I)" + from VSince Since have map: "set (map v_at vps) = set ([(v_at vp) ..< Suc (LTP_p \ i I)])" + by (auto simp: Let_def) + then have kset: "k \ set ([(v_at vp) ..< Suc (LTP_p \ i I)])" using k_def by auto + then obtain x where x: "x \ set vps" "v_at x = k" using k_def map kset + unfolding set_map set_eq_iff image_iff + by metis + then have "VIO \ v k \" using VSince unfolding Since + by (auto simp: Let_def) + } note * = this + show ?thesis + using VSince * + unfolding Since + by (auto simp: Let_def split: enat.splits if_splits + intro!: SAT_VIO.VSince[of _ i "v_at vp"]) + qed (auto intro: SAT_VIO.intros) +next + case (VUntil i vps vp) + then show ?case + proof (cases \) + case (Until \ I \) + {fix k + assume k_def: "k \ v_at vp \ k \ i \ k \ ETP \ (\ \ i + left I)" + from VUntil Until have map: "set (map v_at vps) = set [(ETP_f \ i I) ..< Suc (v_at vp)]" + by (auto simp: Let_def) + then have kset: "k \ set ([(ETP_f \ i I) ..< Suc (v_at vp)])" using k_def by auto + then obtain x where x: "x \ set vps" "v_at x = k" using k_def map kset + unfolding set_map set_eq_iff image_iff + by metis + then have "VIO \ v k \" using VUntil unfolding Until + by (auto simp: Let_def) + } note * = this + then show ?thesis + using VUntil + unfolding Until + by (auto simp: Let_def split: enat.splits if_splits + intro!: SAT_VIO.VUntil) + qed(auto intro: SAT_VIO.intros) +next + case (VSinceInf i li vps) + then show ?case + proof (cases \) + case (Since \ I \) + {fix k + define j where j_def: "j \ case right I of \ \ 0 | enat n \ ETP \ (\ \ i - n)" + assume k_def: "k \ j \ k \ i \ k \ LTP \ (\ \ i - left I)" + from VSinceInf Since j_def have map: "set (map v_at vps) = set [j ..< Suc (LTP_p \ i I)]" + by (auto simp: Let_def) + then have kset: "k \ set ([j ..< Suc (LTP_p \ i I)])" using j_def k_def by auto + then obtain x where x: "x \ set vps" "v_at x = k" using k_def map + unfolding set_map set_eq_iff image_iff + by metis + then have "VIO \ v k \" using VSinceInf unfolding Since + by (auto simp: Let_def) + } note * = this + show ?thesis + using VSinceInf * + unfolding Since + by (auto simp: Let_def intro!: SAT_VIO.VSinceInf) + qed (auto intro: SAT_VIO.intros) +next + case (VUntilInf i hi vps) + then show ?case + proof (cases \) + case (Until \ I \) + obtain n where n_def: "right I = enat n" + using VUntilInf + by (auto simp: Until split: enat.splits) + {fix k + define j where j_def: "j \ LTP \ (\ \ i + n)" + assume k_def: "k \ j \ k \ i \ k \ ETP \ (\ \ i + left I)" + from VUntilInf Until j_def have map: "set (map v_at vps) = set [(ETP_f \ i I) ..< Suc j]" + by (auto simp: Let_def n_def) + then have kset: "k \ set ([(ETP_f \ i I) ..< Suc j])" using k_def j_def by auto + then obtain x where x: "x \ set vps" "v_at x = k" using k_def map + unfolding set_map set_eq_iff image_iff + by metis + then have "VIO \ v k \" using VUntilInf unfolding Until + by (auto simp: Let_def n_def) + } note * = this + then show ?thesis + using VUntilInf + unfolding Until + by (auto simp: Let_def n_def intro: SAT_VIO.VUntilInf split: if_splits enat.splits) + qed(auto intro: SAT_VIO.intros) +qed + +definition "compatible X vs v \ (\x\X. v x \ vs x)" + +definition "compatible_vals X vs = {v. \x \ X. v x \ vs x}" + +lemma compatible_alt: + "compatible X vs v \ v \ compatible_vals X vs" + by (auto simp: compatible_def compatible_vals_def) + +lemma compatible_empty_iff: "compatible {} vs v \ True" + by (auto simp: compatible_def) + +lemma compatible_vals_empty_eq: "compatible_vals {} vs = UNIV" + by (auto simp: compatible_vals_def) + +lemma compatible_union_iff: + "compatible (X \ Y) vs v \ compatible X vs v \ compatible Y vs v" + by (auto simp: compatible_def) + +lemma compatible_vals_union_eq: + "compatible_vals (X \ Y) vs = compatible_vals X vs \ compatible_vals Y vs" + by (auto simp: compatible_vals_def) + +lemma compatible_antimono: + "compatible X vs v \ Y \ X \ compatible Y vs v" + by (auto simp: compatible_def) + +lemma compatible_vals_antimono: + "Y \ X \ compatible_vals X vs \ compatible_vals Y vs" + by (auto simp: compatible_vals_def) + +lemma compatible_extensible: + "(\x. vs x \ {}) \ compatible X vs v \ X \ Y \ \v'. compatible Y vs v' \ (\x\X. v x = v' x)" + using some_in_eq[of "vs _"] by (auto simp: override_on_def compatible_def + intro: exI[where x="override_on v (\x. SOME y. y \ vs x) (Y-X)"]) + +lemmas compatible_vals_extensible = compatible_extensible[unfolded compatible_alt] + +primrec mk_values :: "(('n, 'd) trm \ 'a set) list \ 'a list set" + where "mk_values [] = {[]}" + | "mk_values (T # Ts) = (case T of + (\<^bold>v x, X) \ + let terms = map fst Ts in + if \<^bold>v x \ set terms then + let fst_pos = hd (positions terms (\<^bold>v x)) in (\xs. (xs ! fst_pos) # xs) ` (mk_values Ts) + else set_Cons X (mk_values Ts) + | (\<^bold>c a, X) \ set_Cons X (mk_values Ts))" + +lemma mk_values_nempty: + "{} \ set (map snd tXs) \ mk_values tXs \ {}" + by (induct tXs) + (auto simp: set_Cons_def image_iff split: trm.splits if_splits) + +lemma mk_values_not_Nil: + "{} \ set (map snd tXs) \ tXs \ [] \ vs \ mk_values tXs \ vs \ []" + by (induct tXs) + (auto simp: set_Cons_def image_iff split: trm.splits if_splits) + +lemma mk_values_nth_cong: "\<^bold>v x \ set (map fst tXs) \ + n \ set (positions (map fst tXs) (\<^bold>v x)) \ + m \ set (positions (map fst tXs) (\<^bold>v x)) \ + vs \ mk_values tXs \ + vs ! n = vs ! m" +proof (induct tXs arbitrary: n m vs x) + case (Cons tX tXs) + show ?case + proof (cases n) + case 0 + then show ?thesis + proof (cases m) + case (Suc m') + with 0 show ?thesis + using Cons(2-) Cons.hyps(1)[of x m' _ "tl vs"] positions_eq_nil_iff[of "map fst tXs" "trm.Var x"] + by (fastforce split: if_splits simp: in_set_conv_nth + Let_def nth_Cons' gr0_conv_Suc neq_Nil_conv) + qed simp + next + case n: (Suc n') + then show ?thesis + proof (cases m) + case 0 + with n show ?thesis + using Cons(2-) Cons.hyps(1)[of x _ n' "tl vs"] positions_eq_nil_iff[of "map fst tXs" "trm.Var x"] + by (fastforce split: if_splits simp: in_set_conv_nth + Let_def nth_Cons' gr0_conv_Suc neq_Nil_conv) + next + case (Suc m') + with n show ?thesis + using Cons(1)[of x n' m' "tl vs"] Cons(2-) + by (fastforce simp: set_Cons_def set_positions_eq split: trm.splits if_splits) + qed + qed +qed simp + +definition "mk_values_subset p tXs X + \ (let (fintXs, inftXs) = partition (\tX. finite (snd tX)) tXs in + if inftXs = [] then {p} \ mk_values tXs \ X + else let inf_dups = filter (\tX. (fst tX) \ set (map fst fintXs)) inftXs in + if inf_dups = [] then (if finite X then False else Code.abort STR ''subset on infinite subset'' (\_. {p} \ mk_values tXs \ X)) + else if list_all (\tX. Max (set (positions tXs tX)) < Max (set (positions (map fst tXs) (fst tX)))) inf_dups + then {p} \ mk_values tXs \ X + else (if finite X then False else Code.abort STR ''subset on infinite subset'' (\_. {p} \ mk_values tXs \ X)))" + +lemma mk_values_nemptyI: "\tX \ set tXs. snd tX \ {} \ mk_values tXs \ {}" + by (induct tXs) + (auto simp: Let_def set_Cons_eq split: prod.splits trm.splits) + +lemma infinite_mk_values1: "\tX \ set tXs. snd tX \ {} \ tY \ set tXs \ + \Y. (fst tY, Y) \ set tXs \ infinite Y \ infinite (mk_values tXs)" +proof (induct tXs arbitrary: tY) + case (Cons tX tXs) + show ?case + unfolding Let_def image_iff mk_values.simps split_beta + trm.split[of infinite] if_split[of infinite] + proof (safe, goal_cases var_in var_out const) + case (var_in x) + hence "\tX\set tXs. snd tX \ {}" + by (simp add: Cons.prems(1)) + moreover have "\Z. (trm.Var x, Z) \ set tXs \ infinite Z" + using Cons.prems(2,3) var_in + by (cases "tY \ set tXs"; clarsimp) + (metis (no_types, lifting) Cons.hyps Cons.prems(1) + finite_imageD inj_on_def list.inject list.set_intros(2)) + ultimately have "infinite (mk_values tXs)" + using Cons.hyps var_in + by auto + moreover have "inj (\xs. xs ! hd (positions (map fst tXs) (trm.Var x)) # xs)" + by (clarsimp simp: inj_on_def) + ultimately show ?case + using var_in(3) finite_imageD inj_on_subset + by fastforce + next + case (var_out x) + hence "infinite (snd tX)" + using Cons + by (metis infinite_set_ConsI(2) insert_iff list.simps(15) prod.collapse) + moreover have "mk_values tXs \ {}" + using Cons.prems + by (auto intro!: mk_values_nemptyI) + then show ?case + using Cons var_out infinite_set_ConsI(1)[OF \mk_values tXs \ {}\ \infinite (snd tX)\] + by auto + next + case (const c) + hence "infinite (snd tX)" + using Cons + by (metis infinite_set_ConsI(2) insert_iff list.simps(15) prod.collapse) + moreover have "mk_values tXs \ {}" + using Cons.prems + by (auto intro!: mk_values_nemptyI) + then show ?case + using Cons const infinite_set_ConsI(1)[OF \mk_values tXs \ {}\ \infinite (snd tX)\] + by auto + qed +qed simp + +lemma infinite_mk_values2: "\tX\set tXs. snd tX \ {} \ + tY \ set tXs \ infinite (snd tY) \ + Max (set (positions tXs tY)) \ Max (set (positions (map fst tXs) (fst tY))) \ + infinite (mk_values tXs)" +proof (induct tXs arbitrary: tY) + case (Cons tX tXs) + hence obs1: "\tX\set tXs. snd tX \ {}" + by (simp add: Cons.prems(1)) + note IH = Cons.hyps[OF obs1 _ \infinite (snd tY)\] + have obs2: "tY \ set tXs \ + Max (set (positions (map fst tXs) (fst tY))) \ Max (set (positions tXs tY))" + using Cons.prems(4) unfolding list.map + by (metis Max_set_positions_Cons_tl Suc_le_mono positions_eq_nil_iff set_empty2 subset_empty subset_positions_map_fst) + show ?case + unfolding Let_def image_iff mk_values.simps split_beta + trm.split[of infinite] if_split[of infinite] + proof (safe, goal_cases var_in var_out const) + case (var_in x) + then show ?case + proof (cases "tY \ set tXs") + case True + hence "infinite ((\Xs. Xs ! hd (positions (map fst tXs) (trm.Var x)) # Xs) ` mk_values tXs)" + using IH[OF True obs2[OF True]] finite_imageD inj_on_def by blast + then show "False" + using var_in by blast + next + case False + have "Max (set (positions (map fst (tX # tXs)) (fst tY))) + = Suc (Max (set (positions (map fst tXs) (fst tY))))" + using Cons.prems var_in + by (simp only: list.map(2)) + (subst Max_set_positions_Cons_tl; force simp: image_iff) + moreover have "tY \ set tXs \ Max (set (positions (tX # tXs) tY)) = (0::nat)" + using Cons.prems Max_set_positions_Cons_hd by fastforce + ultimately show "False" + using Cons.prems(4) False + by linarith + qed + next + case (var_out x) + then show ?case + proof (cases "tY \ set tXs") + case True + hence "infinite (mk_values tXs)" + using IH obs2 by blast + hence "infinite (set_Cons (snd tX) (mk_values tXs))" + by (metis Cons.prems(1) infinite_set_ConsI(2) list.set_intros(1)) + then show "False" + using var_out by blast + next + case False + hence "snd tY = snd tX" and "infinite (snd tX)" + using var_out Cons.prems + by auto + hence "infinite (set_Cons (snd tX) (mk_values tXs))" + by (simp add: infinite_set_ConsI(1) mk_values_nemptyI obs1) + then show "False" + using var_out by blast + qed + next + case (const c) + then show ?case + proof (cases "tY \ set tXs") + case True + hence "infinite (mk_values tXs)" + using IH obs2 by blast + hence "infinite (set_Cons (snd tX) (mk_values tXs))" + by (metis Cons.prems(1) infinite_set_ConsI(2) list.set_intros(1)) + then show "False" + using const by blast + next + case False + hence "infinite (set_Cons (snd tX) (mk_values tXs))" + using const Cons.prems + by (simp add: infinite_set_ConsI(1) mk_values_nemptyI obs1) + then show "False" + using const by blast + qed + qed +qed simp + +lemma mk_values_subset_iff: "\tX \ set tXs. snd tX \ {} \ + mk_values_subset p tXs X \ {p} \ mk_values tXs \ X" + unfolding mk_values_subset_def image_iff Let_def comp_def split_beta if_split_eq1 + partition_filter1 partition_filter2 o_def set_map set_filter filter_filter bex_simps +proof safe + assume "\tX\set tXs. snd tX \ {}" and "finite X" + and filter1: "filter (\xy. infinite (snd xy) \ (\ab. (ab \ set tXs \ finite (snd ab)) \ fst xy = fst ab)) tXs = []" + and filter2: "filter (\x. infinite (snd x)) tXs \ []" + then obtain tY where "tY \ set tXs" and "infinite (snd tY)" + by (meson filter_False) + moreover have "\Y. (fst tY, Y) \ set tXs \ infinite Y" + using filter1 calculation + by (auto simp: filter_empty_conv) + ultimately have "infinite (mk_values tXs)" + using infinite_mk_values1[OF \\tX\set tXs. snd tX \ {}\] + by auto + hence "infinite ({p} \ mk_values tXs)" + using finite_cartesian_productD2 by auto + thus "{p} \ mk_values tXs \ X \ False" + using \finite X\ + by (simp add: finite_subset) +next + assume "\tX\set tXs. snd tX \ {}" + and "finite X" + and ex_dupl_inf: "\ list_all (\tX. Max (set (positions tXs tX)) + < Max (set (positions (map fst tXs) (fst tX)))) + (filter (\xy. infinite (snd xy) \ (\ab. (ab \ set tXs \ finite (snd ab)) \ fst xy = fst ab)) tXs)" + and filter: "filter (\x. infinite (snd x)) tXs \ []" + then obtain tY and Z where "tY \ set tXs" + and "infinite (snd tY)" + and "(fst tY, Z) \ set tXs" + and "finite Z" + and "Max (set (positions tXs tY)) \ Max (set (positions (map fst tXs) (fst tY)))" + by (auto simp: list_all_iff) + hence "infinite (mk_values tXs)" + using infinite_mk_values2[OF \\tX\set tXs. snd tX \ {}\ \tY \ set tXs\] + by auto + hence "infinite ({p} \ mk_values tXs)" + using finite_cartesian_productD2 by auto + thus "{p} \ mk_values tXs \ X \ False" + using \finite X\ + by (simp add: finite_subset) +qed auto + +lemma mk_values_sound: "cs \ mk_values (vs\<^bold>\ts\<^bold>\) \ + \v\compatible_vals (fv (r \ ts)) vs. cs = v\<^bold>\ts\<^bold>\" +proof (induct ts arbitrary: cs vs) + case (Cons t ts) + show ?case + proof(cases t) + case (Var x) + let ?Ts = "vs\<^bold>\ts\<^bold>\" + have "vs\<^bold>\(t # ts)\<^bold>\ = (\<^bold>v x, vs x) # ?Ts" + using Var by (simp add: eval_trms_set_def) + show ?thesis + proof (cases "\<^bold>v x \ set ts") + case True + then obtain n where n_in: "n \ set (positions ts (\<^bold>v x))" + and nth_n: "ts ! n = \<^bold>v x" + by (meson fst_pos_in_positions nth_fst_pos) + hence n_in': "n \ set (positions (map fst ?Ts) (\<^bold>v x))" + by (induct ts arbitrary: n) + (auto simp: eval_trms_set_def split: if_splits) + have key: "\<^bold>v x \ set (map fst ?Ts)" + using True by (induct ts) + (auto simp: eval_trms_set_def) + then obtain a as + where as_head: "as ! (hd (positions (map fst ?Ts) (\<^bold>v x))) = a" + and as_tail: "as \ mk_values ?Ts" + and as_shape: "cs = a # as" + using Cons(2) + by (clarsimp simp add: eval_trms_set_def Var image_iff) + obtain v where v_hyps: "v \ compatible_vals (fv (r \ ts)) vs" + "as = v\<^bold>\ts\<^bold>\" + using Cons(1)[OF as_tail] by blast + hence as'_nth: "as ! n = v x" + using nth_n positions_length[OF n_in] + by (simp add: eval_trms_def) + have evals_neq_Nil: "?Ts \ []" + using key by auto + moreover have "positions (map fst ?Ts) (\<^bold>v x) \ []" + using positions_eq_nil_iff[of "map fst ?Ts" "\<^bold>v x"] key + by fastforce + ultimately have as_hyp: "a = as ! n" + using mk_values_nth_cong[OF key hd_in_set n_in' as_tail] as_head by blast + thus ?thesis + using Var as_shape True v_hyps as'_nth + by (auto simp: compatible_vals_def eval_trms_def intro!: exI[of _ v]) + next + case False + hence *: "\<^bold>v x \ set (map fst ?Ts)" + proof (induct ts arbitrary: x) + case (Cons a ts) + then show ?case + by (cases a) (auto simp: eval_trms_set_def) + qed (simp add: eval_trms_set_def) + from Cons(2) False show ?thesis + unfolding set_Cons_def eval_trms_set_def Let_def list.simps Var + *[THEN eq_False[THEN iffD2], unfolded eval_trms_set_def] if_False + mk_values.simps eval_trm_set.simps prod.case trm.case mem_Collect_eq + proof (elim exE conjE, goal_cases) + case (1 a as) + with Cons(1)[of as vs] obtain v where "v \ compatible_vals (fv (r \ ts)) vs" "as = v\<^bold>\ts\<^bold>\" + by (auto simp: eval_trms_set_def) + with 1 show ?case + by (auto simp: eval_trms_set_def eval_trms_def compatible_vals_def in_fv_trm_conv + intro!: exI[of _ "v(x := a)"] eval_trm_fv_cong) + qed + qed + next + case (Const c) + then show ?thesis + using Cons(1)[of _ vs] Cons(2) + by (auto simp: eval_trms_set_def set_Cons_def + eval_trms_def compatible_def) + qed +qed (simp add: eval_trms_set_def eval_trms_def compatible_vals_def) + +lemma fst_eval_trm_set[simp]: + "fst (vs\t\) = t" + by (cases t; clarsimp) + +lemma mk_values_complete: "cs = v\<^bold>\ts\<^bold>\ \ + v \ compatible_vals (fv (r \ ts)) vs \ + cs \ mk_values (vs\<^bold>\ts\<^bold>\)" +proof (induct ts arbitrary: v cs vs) + case (Cons t ts) + then obtain a as + where a_def: "a = v\t\" + and as_def: "as = v\<^bold>\ts\<^bold>\" + and cs_cons: "cs = a # as" + by (auto simp: eval_trms_def) + have compat_v_vs: "v \ compatible_vals (fv (r \ ts)) vs" + using Cons.prems + by (auto simp: compatible_vals_def) + hence mk_values_ts: "as \ mk_values (vs\<^bold>\ts\<^bold>\)" + using Cons.hyps[OF as_def] + unfolding eval_trms_set_def by blast + show ?case + proof (cases "t") + case (Var x) + then show ?thesis + proof (cases "\<^bold>v x \ set ts") + case True + then obtain n + where n_head: "n = hd (positions ts (\<^bold>v x))" + and n_in: "n \ set (positions ts (\<^bold>v x))" + and nth_n: "ts ! n = \<^bold>v x" + by (simp_all add: hd_positions_eq_fst_pos nth_fst_pos fst_pos_in_positions) + hence n_in': "n = hd (positions (map fst (vs\<^bold>\ts\<^bold>\)) (\<^bold>v x))" + by (clarsimp simp: eval_trms_set_def o_def) + moreover have "as ! n = a" + using a_def as_def nth_n Var n_in True positions_length + by (fastforce simp: eval_trms_def) + moreover have "\<^bold>v x \ set (map fst (vs\<^bold>\ts\<^bold>\))" + using True by (induct ts) + (auto simp: eval_trms_set_def) + ultimately show ?thesis + using mk_values_ts cs_cons + by (clarsimp simp: eval_trms_set_def Var image_iff) + next + case False + then show ?thesis + using Var cs_cons mk_values_ts Cons.prems a_def + by (clarsimp simp: eval_trms_set_def image_iff + set_Cons_def compatible_vals_def split: trm.splits) + qed + next + case (Const a) + then show ?thesis + using cs_cons mk_values_ts Cons.prems a_def + by (clarsimp simp: eval_trms_set_def image_iff + set_Cons_def compatible_vals_def split: trm.splits) + qed +qed (simp add: compatible_vals_def + eval_trms_set_def eval_trms_def) + +definition "mk_values_subset_Compl r vs ts i = ({r} \ mk_values (vs\<^bold>\ts\<^bold>\) \ - \ \ i)" + +fun check_values where + "check_values _ _ _ None = None" +| "check_values vs (\<^bold>c c # ts) (u # us) f = (if c = u then check_values vs ts us f else None)" +| "check_values vs (\<^bold>v x # ts) (u # us) (Some v) = (if u \ vs x \ (v x = Some u \ v x = None) then check_values vs ts us (Some (v(x \ u))) else None)" +| "check_values vs [] [] f = f" +| "check_values _ _ _ _ = None" + +lemma mk_values_alt: + "mk_values (vs\<^bold>\ts\<^bold>\) = + {cs. \v\compatible_vals (\(fv_trm ` set ts)) vs. cs = v\<^bold>\ts\<^bold>\}" + by (auto dest!: mk_values_sound intro: mk_values_complete) + +lemma check_values_neq_NoneI: + assumes "v \ compatible_vals (\ (fv_trm ` set ts) - dom f) vs" "\x y. f x = Some y \ y \ vs x" + shows "check_values vs ts ((\x. case f x of None \ v x | Some y \ y)\<^bold>\ts\<^bold>\) (Some f) \ None" + using assms +proof (induct ts arbitrary: f) + case (Cons t ts) + then show ?case + proof (cases t) + case (Var x) + show ?thesis + proof (cases "f x") + case None + with Cons(2) Var have v_in[simp]: "v x \ vs x" + by (auto simp: compatible_vals_def) + from Cons(2) have "v \ compatible_vals (\ (fv_trm ` set ts) - dom (f(x \ v x))) vs" + by (auto simp: compatible_vals_def) + then have "check_values vs ts + ((\z. case (f(x \ v x)) z of None \ v z | Some y \ y)\<^bold>\ts\<^bold>\) + (Some (f(x \ v x))) \ None" + using Cons(3) None Var + by (intro Cons(1)) (auto simp: compatible_vals_def split: if_splits) + then show ?thesis + by (elim eq_neq_eq_imp_neq[OF _ _ refl, rotated]) + (auto simp add: eval_trms_def compatible_vals_def Var None split: if_splits option.splits + intro!: arg_cong2[of _ _ _ _ "check_values vs ts"] eval_trm_fv_cong) + next + case (Some y) + with Cons(1)[of f] Cons(2-) Var fun_upd_triv[of f x] show ?thesis + by (auto simp: domI eval_trms_def compatible_vals_def split: option.splits) + qed + next + case (Const c) + with Cons show ?thesis + by (auto simp: eval_trms_def compatible_vals_def split: option.splits) + qed +qed (simp add: eval_trms_def) + +lemma check_values_eq_NoneI: + "\v\compatible_vals (\ (fv_trm ` set ts) - dom f) vs. us \ (\x. case f x of None \ v x | Some y \ y)\<^bold>\ts\<^bold>\ \ + check_values vs ts us (Some f) = None" +proof (induct vs ts us "Some f" arbitrary: f rule: check_values.induct) + case (3 vs x ts u us v) + show ?case + unfolding check_values.simps if_split_eq1 simp_thms + proof (intro impI 3(1), safe, goal_cases) + case (1 v') + with 3(2) show ?case + by (auto simp: insert_dom domI eval_trms_def intro!: eval_trm_fv_cong split: if_splits dest!: bspec[of _ _ "v'"]) + next + case (2 v') + with 3(2) show ?case + by (auto simp: insert_dom domI compatible_vals_def eval_trms_def intro!: eval_trm_fv_cong split: if_splits option.splits dest!: spec[of _ "v'(x := u)"]) + qed +qed (auto simp: compatible_vals_def eval_trms_def) + +lemma mk_values_subset_Compl_code[code]: + "mk_values_subset_Compl r vs ts i = (\(q, us) \ \ \ i. q \ r \ check_values vs ts us (Some Map.empty) = None)" + unfolding mk_values_subset_Compl_def eval_trms_set_def[symmetric] mk_values_alt +proof (safe, goal_cases) + case (1 _ us y) + then show ?case + by (auto simp: subset_eq check_values_eq_NoneI[where f=Map.empty, simplified] dest!: spec[of _ us]) +qed (auto simp: subset_eq dest!: check_values_neq_NoneI[where f=Map.empty, simplified]) + +subsection \Executable Variant of the Checker\ + +fun s_check_exec :: "('n, 'd) envset \ ('n, 'd) formula \ ('n, 'd) sproof \ bool" +and v_check_exec :: "('n, 'd) envset \ ('n, 'd) formula \ ('n, 'd) vproof \ bool" where + "s_check_exec vs f p = (case (f, p) of + (\, STT i) \ True + | (r \ ts, SPred i s ts') \ + (r = s \ ts = ts' \ mk_values_subset r (vs\<^bold>\ts\<^bold>\) (\ \ i)) + | (x \<^bold>\ c, SEq_Const i x' c') \ + (c = c' \ x = x' \ vs x = {c}) + | (\\<^sub>F \, SNeg vp) \ v_check_exec vs \ vp + | (\ \\<^sub>F \, SOrL sp1) \ s_check_exec vs \ sp1 + | (\ \\<^sub>F \, SOrR sp2) \ s_check_exec vs \ sp2 + | (\ \\<^sub>F \, SAnd sp1 sp2) \ s_check_exec vs \ sp1 \ s_check_exec vs \ sp2 \ s_at sp1 = s_at sp2 + | (\ \\<^sub>F \, SImpL vp1) \ v_check_exec vs \ vp1 + | (\ \\<^sub>F \, SImpR sp2) \ s_check_exec vs \ sp2 + | (\ \\<^sub>F \, SIffSS sp1 sp2) \ s_check_exec vs \ sp1 \ s_check_exec vs \ sp2 \ s_at sp1 = s_at sp2 + | (\ \\<^sub>F \, SIffVV vp1 vp2) \ v_check_exec vs \ vp1 \ v_check_exec vs \ vp2 \ v_at vp1 = v_at vp2 + | (\\<^sub>Fx. \, SExists y val sp) \ (x = y \ s_check_exec (vs (x := {val})) \ sp) + | (\\<^sub>Fx. \, SForall y sp_part) \ (let i = s_at (part_hd sp_part) + in x = y \ (\(sub, sp) \ SubsVals sp_part. s_at sp = i \ s_check_exec (vs (x := sub)) \ sp)) + | (\<^bold>Y I \, SPrev sp) \ + (let j = s_at sp; i = s_at (SPrev sp) in + i = j+1 \ mem (\ \ i) I \ s_check_exec vs \ sp) + | (\<^bold>X I \, SNext sp) \ + (let j = s_at sp; i = s_at (SNext sp) in + j = i+1 \ mem (\ \ j) I \ s_check_exec vs \ sp) + | (\<^bold>P I \, SOnce i sp) \ + (let j = s_at sp in + j \ i \ mem (\ \ i - \ \ j) I \ s_check_exec vs \ sp) + | (\<^bold>F I \, SEventually i sp) \ + (let j = s_at sp in + j \ i \ mem (\ \ j - \ \ i) I \ s_check_exec vs \ sp) + | (\<^bold>H I \, SHistoricallyOut i) \ + \ \ i < \ \ 0 + left I + | (\<^bold>H I \, SHistorically i li sps) \ + (li = (case right I of \ \ 0 | enat b \ ETP \ (\ \ i - b)) + \ \ \ 0 + left I \ \ \ i + \ map s_at sps = [li ..< (LTP_p \ i I) + 1] + \ (\sp \ set sps. s_check_exec vs \ sp)) + | (\<^bold>G I \, SAlways i hi sps) \ + (hi = (case right I of enat b \ LTP_f \ i b) + \ right I \ \ + \ map s_at sps = [(ETP_f \ i I) ..< hi + 1] + \ (\sp \ set sps. s_check_exec vs \ sp)) + | (\ \<^bold>S I \, SSince sp2 sp1s) \ + (let i = s_at (SSince sp2 sp1s); j = s_at sp2 in + j \ i \ mem (\ \ i - \ \ j) I + \ map s_at sp1s = [j+1 ..< i+1] + \ s_check_exec vs \ sp2 + \ (\sp1 \ set sp1s. s_check_exec vs \ sp1)) + | (\ \<^bold>U I \, SUntil sp1s sp2) \ + (let i = s_at (SUntil sp1s sp2); j = s_at sp2 in + j \ i \ mem (\ \ j - \ \ i) I + \ map s_at sp1s = [i ..< j] \ s_check_exec vs \ sp2 + \ (\sp1 \ set sp1s. s_check_exec vs \ sp1)) + | ( _ , _) \ False)" +| "v_check_exec vs f p = (case (f, p) of + (\, VFF i) \ True + | (r \ ts, VPred i pred ts') \ + (r = pred \ ts = ts' \ mk_values_subset_Compl r vs ts i) + | (x \<^bold>\ c, VEq_Const i x' c') \ + (c = c' \ x = x' \ c \ vs x) + | (\\<^sub>F \, VNeg sp) \ s_check_exec vs \ sp + | (\ \\<^sub>F \, VOr vp1 vp2) \ v_check_exec vs \ vp1 \ v_check_exec vs \ vp2 \ v_at vp1 = v_at vp2 + | (\ \\<^sub>F \, VAndL vp1) \ v_check_exec vs \ vp1 + | (\ \\<^sub>F \, VAndR vp2) \ v_check_exec vs \ vp2 + | (\ \\<^sub>F \, VImp sp1 vp2) \ s_check_exec vs \ sp1 \ v_check_exec vs \ vp2 \ s_at sp1 = v_at vp2 + | (\ \\<^sub>F \, VIffSV sp1 vp2) \ s_check_exec vs \ sp1 \ v_check_exec vs \ vp2 \ s_at sp1 = v_at vp2 + | (\ \\<^sub>F \, VIffVS vp1 sp2) \ v_check_exec vs \ vp1 \ s_check_exec vs \ sp2 \ v_at vp1 = s_at sp2 + | (\\<^sub>Fx. \, VExists y vp_part) \ (let i = v_at (part_hd vp_part) + in x = y \ (\(sub, vp) \ SubsVals vp_part. v_at vp = i \ v_check_exec (vs (x := sub)) \ vp)) + | (\\<^sub>Fx. \, VForall y val vp) \ (x = y \ v_check_exec (vs (x := {val})) \ vp) + | (\<^bold>Y I \, VPrev vp) \ + (let j = v_at vp; i = v_at (VPrev vp) in + i = j+1 \ v_check_exec vs \ vp) + | (\<^bold>Y I \, VPrevZ) \ True + | (\<^bold>Y I \, VPrevOutL i) \ + i > 0 \ \ \ i < left I + | (\<^bold>Y I \, VPrevOutR i) \ + i > 0 \ enat (\ \ i) > right I + | (\<^bold>X I \, VNext vp) \ + (let j = v_at vp; i = v_at (VNext vp) in + j = i+1 \ v_check_exec vs \ vp) + | (\<^bold>X I \, VNextOutL i) \ + \ \ (i+1) < left I + | (\<^bold>X I \, VNextOutR i) \ + enat (\ \ (i+1)) > right I + | (\<^bold>P I \, VOnceOut i) \ + \ \ i < \ \ 0 + left I + | (\<^bold>P I \, VOnce i li vps) \ + (li = (case right I of \ \ 0 | enat b \ ETP_p \ i b) + \ \ \ 0 + left I \ \ \ i + \ map v_at vps = [li ..< (LTP_p \ i I) + 1] + \ (\vp \ set vps. v_check_exec vs \ vp)) + | (\<^bold>F I \, VEventually i hi vps) \ + (hi = (case right I of enat b \ LTP_f \ i b) \ right I \ \ + \ map v_at vps = [(ETP_f \ i I) ..< hi + 1] + \ (\vp \ set vps. v_check_exec vs \ vp)) + | (\<^bold>H I \, VHistorically i vp) \ + (let j = v_at vp in + j \ i \ mem (\ \ i - \ \ j) I \ v_check_exec vs \ vp) + | (\<^bold>G I \, VAlways i vp) \ + (let j = v_at vp + in j \ i \ mem (\ \ j - \ \ i) I \ v_check_exec vs \ vp) + | (\ \<^bold>S I \, VSinceOut i) \ + \ \ i < \ \ 0 + left I + | (\ \<^bold>S I \, VSince i vp1 vp2s) \ + (let j = v_at vp1 in + (case right I of \ \ True | enat b \ ETP_p \ i b \ j) \ j \ i + \ \ \ 0 + left I \ \ \ i + \ map v_at vp2s = [j ..< (LTP_p \ i I) + 1] \ v_check_exec vs \ vp1 + \ (\vp2 \ set vp2s. v_check_exec vs \ vp2)) + | (\ \<^bold>S I \, VSinceInf i li vp2s) \ + (li = (case right I of \ \ 0 | enat b \ ETP_p \ i b) + \ \ \ 0 + left I \ \ \ i + \ map v_at vp2s = [li ..< (LTP_p \ i I) + 1] + \ (\vp2 \ set vp2s. v_check_exec vs \ vp2)) + | (\ \<^bold>U I \, VUntil i vp2s vp1) \ + (let j = v_at vp1 in + (case right I of \ \ True | enat b \ j < LTP_f \ i b) \ i \ j + \ map v_at vp2s = [ETP_f \ i I ..< j + 1] \ v_check_exec vs \ vp1 + \ (\vp2 \ set vp2s. v_check_exec vs \ vp2)) + | (\ \<^bold>U I \, VUntilInf i hi vp2s) \ + (hi = (case right I of enat b \ LTP_f \ i b) \ right I \ \ + \ map v_at vp2s = [ETP_f \ i I ..< hi + 1] + \ (\vp2 \ set vp2s. v_check_exec vs \ vp2)) + | ( _ , _) \ False)" + +declare s_check_exec.simps[simp del] v_check_exec.simps[simp del] +simps_of_case s_check_exec_simps[simp, code]: s_check_exec.simps[unfolded prod.case] (splits: formula.split sproof.split) +simps_of_case v_check_exec_simps[simp, code]: v_check_exec.simps[unfolded prod.case] (splits: formula.split vproof.split) + + +lemma check_fv_cong: + assumes "\x \ fv \. v x = v' x" + shows "s_check v \ sp \ s_check v' \ sp" "v_check v \ vp \ v_check v' \ vp" + using assms +proof (induct \ arbitrary: v v' sp vp) + case TT + { + case 1 + then show ?case + by (cases sp) auto + next + case 2 + then show ?case + by (cases vp) auto + } +next + case FF + { + case 1 + then show ?case + by (cases sp) auto + next + case 2 + then show ?case + by (cases vp) auto + } +next + case (Pred p ts) + { + case 1 + with Pred show ?case using eval_trms_fv_cong[of ts v v'] + by (cases sp) auto + next + case 2 + with Pred show ?case using eval_trms_fv_cong[of ts v v'] + by (cases vp) auto + } + case (Eq_Const x c) + { + case 1 + then show ?case + by (cases sp) auto + next + case 2 + then show ?case + by (cases vp) auto + } +next + case (Neg \) + { + case 1 + with Neg[of v v'] show ?case + by (cases sp) auto + next + case 2 + with Neg[of v v'] show ?case + by (cases vp) auto + } +next + case (Or \1 \2) + { + case 1 + with Or[of v v'] show ?case + by (cases sp) auto + next + case 2 + with Or[of v v'] show ?case + by (cases vp) auto + } +next + case (And \1 \2) + { + case 1 + with And[of v v'] show ?case + by (cases sp) auto + next + case 2 + with And[of v v'] show ?case + by (cases vp) auto + } +next + case (Imp \1 \2) + { + case 1 + with Imp[of v v'] show ?case + by (cases sp) auto + next + case 2 + with Imp[of v v'] show ?case + by (cases vp) auto + } +next + case (Iff \1 \2) + { + case 1 + with Iff[of v v'] show ?case + by (cases sp) auto + next + case 2 + with Iff[of v v'] show ?case + by (cases vp) auto + } +next + case (Exists x \) + { + case 1 + with Exists[of "v(x := z)" "v'(x := z)" for z] show ?case + by (cases sp) (auto simp: fun_upd_def) + next + case 2 + with Exists[of "v(x := z)" "v'(x := z)" for z] show ?case + by (cases vp) (auto simp: fun_upd_def) + } +next + case (Forall x \) + { + case 1 + with Forall[of "v(x := z)" "v'(x := z)" for z] show ?case + by (cases sp) (auto simp: fun_upd_def) + next + case 2 + with Forall[of "v(x := z)" "v'(x := z)" for z] show ?case + by (cases vp) (auto simp: fun_upd_def) + } +next + case (Prev I \) + { + case 1 + with Prev[of v v'] show ?case + by (cases sp) auto + next + case 2 + with Prev[of v v'] show ?case + by (cases vp) auto + } +next + case (Next I \) + { + case 1 + with Next[of v v'] show ?case + by (cases sp) auto + next + case 2 + with Next[of v v'] show ?case + by (cases vp) auto + } +next + case (Once I \) + { + case 1 + with Once[of v v'] show ?case + by (cases sp) auto + next + case 2 + with Once[of v v'] show ?case + by (cases vp) auto + } +next + case (Historically I \) + { + case 1 + with Historically[of v v'] show ?case + by (cases sp) auto + next + case 2 + with Historically[of v v'] show ?case + by (cases vp) auto + } +next + case (Eventually I \) + { + case 1 + with Eventually[of v v'] show ?case + by (cases sp) auto + next + case 2 + with Eventually[of v v'] show ?case + by (cases vp) auto + } +next + case (Always I \) + { + case 1 + with Always[of v v'] show ?case + by (cases sp) auto + next + case 2 + with Always[of v v'] show ?case + by (cases vp) auto + } +next + case (Since \1 I \2) + { + case 1 + with Since[of v v'] show ?case + by (cases sp) auto + next + case 2 + with Since[of v v'] show ?case + by (cases vp) auto + } +next + case (Until \1 I \2) + { + case 1 + with Until[of v v'] show ?case + by (cases sp) auto + next + case 2 + with Until[of v v'] show ?case + by (cases vp) auto + } +qed + +lemma s_check_fun_upd_notin[simp]: + "x \ fv \ \ s_check (v(x := t)) \ sp = s_check v \ sp" + by (rule check_fv_cong) auto +lemma v_check_fun_upd_notin[simp]: + "x \ fv \ \ v_check (v(x := t)) \ sp = v_check v \ sp" + by (rule check_fv_cong) auto + +lemma SubsVals_nonempty: "(X, t) \ SubsVals part \ X \ {}" + by transfer (auto simp: partition_on_def image_iff) + +lemma compatible_vals_nonemptyI: "\x. vs x \ {} \ compatible_vals A vs \ {}" + by (auto simp: compatible_vals_def intro!: bchoice) + +lemma compatible_vals_fun_upd: "compatible_vals A (vs(x := X)) = + (if x \ A then {v \ compatible_vals (A - {x}) vs. v x \ X} else compatible_vals A vs)" + unfolding compatible_vals_def + by auto + +lemma fun_upd_in_compatible_vals: "v \ compatible_vals (A - {x}) vs \ v(x := t) \ compatible_vals (A - {x}) vs" + unfolding compatible_vals_def + by auto + +lemma fun_upd_in_compatible_vals_in: "v \ compatible_vals (A - {x}) vs \ t \ vs x \ v(x := t) \ compatible_vals A vs" + unfolding compatible_vals_def + by auto + +lemma fun_upd_in_compatible_vals_notin: "x \ A \ v \ compatible_vals A vs \ v(x := t) \ compatible_vals A vs" + unfolding compatible_vals_def + by auto + +lemma check_exec_check: + assumes "\x. vs x \ {}" + shows "s_check_exec vs \ sp \ (\v \ compatible_vals (fv \) vs. s_check v \ sp)" + and "v_check_exec vs \ vp \ (\v \ compatible_vals (fv \) vs. v_check v \ vp)" + using assms +proof (induct \ arbitrary: vs sp vp) + case TT + { + case 1 + then show ?case using compatible_vals_nonemptyI + by (cases sp) + auto + next + case 2 + then show ?case using compatible_vals_nonemptyI + by auto + } +next + case FF + { + case 1 + then show ?case using compatible_vals_nonemptyI + by (cases sp) + auto + next + case 2 + then show ?case using compatible_vals_nonemptyI + by (cases vp) + auto + } +next + case (Pred p ts) + { + case 1 + have obs: "\tX\set (vs\<^bold>\ts\<^bold>\). snd tX \ {}" + using \\x. vs x \ {}\ + proof (induct ts) + case (Cons a ts) + then show ?case + by (cases a) (auto simp: eval_trms_set_def) + qed (auto simp: eval_trms_set_def) + show ?case + using 1 compatible_vals_nonemptyI[OF 1] + mk_values_complete[OF refl, of _ p ts vs] mk_values_sound[of _ vs ts p] + by (cases sp) + (auto 6 0 simp: mk_values_subset_iff[OF obs] simp del: fv.simps) + next + case 2 + then show ?case using compatible_vals_nonemptyI[OF 2] + mk_values_complete[OF refl, of _ p ts vs] mk_values_sound[of _ vs ts p] + by (cases vp) + (auto 6 0 simp: mk_values_subset_Compl_def eval_trms_set_def simp del: fv.simps) + } +next + case (Eq_Const x c) + { + case 1 + then show ?case + by (cases sp) (auto simp: compatible_vals_def) + next + case 2 + then show ?case + by (cases vp) (auto simp: compatible_vals_def) + } +next + case (Neg \) + { + case 1 + then show ?case + using Neg.hyps(2) compatible_vals_nonemptyI[OF 1] + by (cases sp) auto + next + case 2 + then show ?case + using Neg.hyps(1) compatible_vals_nonemptyI[OF 2] + by (cases vp) auto + } +next + case (Or \1 \2) + { + case 1 + with compatible_vals_nonemptyI[OF 1, of "fv \1 \ fv \2"] show ?case + proof (cases sp) + case (SOrL sp') + from check_fv_cong(1)[of \1 _ _ sp'] show ?thesis + unfolding SOrL s_check_exec_simps s_check_simps fv.simps Or(1)[OF 1, of sp'] + by (metis (mono_tags, lifting) 1 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + next + case (SOrR sp') + from check_fv_cong(1)[of \2 _ _ sp'] show ?thesis + unfolding SOrR s_check_exec_simps s_check_simps fv.simps Or(3)[OF 1, of sp'] + by (metis (mono_tags, lifting) 1 IntE Un_upper2 compatible_vals_extensible compatible_vals_union_eq) + qed (auto simp: compatible_vals_union_eq) + next + case 2 + with compatible_vals_nonemptyI[OF 2, of "fv \1 \ fv \2"] show ?case + proof (cases vp) + case (VOr vp1 vp2) + from check_fv_cong(2)[of \1 _ _ vp1] check_fv_cong(2)[of \2 _ _ vp2] show ?thesis + unfolding VOr v_check_exec_simps v_check_simps fv.simps ball_conj_distrib + Or(2)[OF 2, of vp1] Or(4)[OF 2, of vp2] + ball_triv_nonempty[OF compatible_vals_nonemptyI[OF 2, of "fv \1 \ fv \2"]] + proof (intro arg_cong2[of _ _ _ _ "(\)"] refl, goal_cases \1 \2) + case \1 + then show ?case + by (metis (mono_tags, lifting) 2 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + next + case \2 + then show ?case + by (metis (mono_tags, lifting) 2 IntE Un_upper2 compatible_vals_extensible compatible_vals_union_eq) + qed + qed (auto simp: compatible_vals_union_eq) + } +next + case (And \1 \2) + { + case 1 + with compatible_vals_nonemptyI[OF 1, of "fv \1 \ fv \2"] show ?case + proof (cases sp) + case (SAnd sp1 sp2) + from check_fv_cong(1)[of \1 _ _ sp1] check_fv_cong(1)[of \2 _ _ sp2] show ?thesis + unfolding SAnd s_check_exec_simps s_check_simps fv.simps ball_conj_distrib + And(1)[OF 1, of sp1] And(3)[OF 1, of sp2] + ball_triv_nonempty[OF compatible_vals_nonemptyI[OF 1, of "fv \1 \ fv \2"]] + proof (intro arg_cong2[of _ _ _ _ "(\)"] refl, goal_cases \1 \2) + case \1 + then show ?case + by (metis (mono_tags, lifting) 1 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + next + case \2 + then show ?case + by (metis (mono_tags, lifting) 1 IntE Un_upper2 compatible_vals_extensible compatible_vals_union_eq) + qed + qed (auto simp: compatible_vals_union_eq) + next + case 2 + with compatible_vals_nonemptyI[OF 2, of "fv \1 \ fv \2"] show ?case + proof (cases vp) + case (VAndL vp') + from check_fv_cong(2)[of \1 _ _ vp'] show ?thesis + unfolding VAndL v_check_exec_simps v_check_simps fv.simps And(2)[OF 2, of vp'] + by (metis (mono_tags, lifting) 2 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + next + case (VAndR vp') + from check_fv_cong(2)[of \2 _ _ vp'] show ?thesis + unfolding VAndR v_check_exec_simps v_check_simps fv.simps And(4)[OF 2, of vp'] + by (metis (mono_tags, lifting) 2 IntE Un_upper2 compatible_vals_extensible compatible_vals_union_eq) + qed (auto simp: compatible_vals_union_eq) + } +next + case (Imp \1 \2) + { + case 1 + with compatible_vals_nonemptyI[OF 1, of "fv \1 \ fv \2"] show ?case + proof (cases sp) + case (SImpL vp') + from check_fv_cong(2)[of \1 _ _ vp'] show ?thesis + unfolding SImpL s_check_exec_simps s_check_simps fv.simps Imp(2)[OF 1, of vp'] + by (metis (mono_tags, lifting) 1 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + next + case (SImpR sp') + from check_fv_cong(1)[of \2 _ _ sp'] show ?thesis + unfolding SImpR s_check_exec_simps s_check_simps fv.simps Imp(3)[OF 1, of sp'] + by (metis (mono_tags, lifting) 1 IntE Un_upper2 compatible_vals_extensible compatible_vals_union_eq) + qed (auto simp: compatible_vals_union_eq) + next + case 2 + with compatible_vals_nonemptyI[OF 2, of "fv \1 \ fv \2"] show ?case + proof (cases vp) + case (VImp sp1 vp2) + from check_fv_cong(1)[of \1 _ _ sp1] check_fv_cong(2)[of \2 _ _ vp2] show ?thesis + unfolding VImp v_check_exec_simps v_check_simps fv.simps ball_conj_distrib + Imp(1)[OF 2, of sp1] Imp(4)[OF 2, of vp2] + ball_triv_nonempty[OF compatible_vals_nonemptyI[OF 2, of "fv \1 \ fv \2"]] + proof (intro arg_cong2[of _ _ _ _ "(\)"] refl, goal_cases \1 \2) + case \1 + then show ?case + by (metis (mono_tags, lifting) 2 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + next + case \2 + then show ?case + by (metis (mono_tags, lifting) 2 IntE Un_upper2 compatible_vals_extensible compatible_vals_union_eq) + qed + qed (auto simp: compatible_vals_union_eq) + } +next + case (Iff \1 \2) + { + case 1 + with compatible_vals_nonemptyI[OF 1, of "fv \1 \ fv \2"] show ?case + proof (cases sp) + case (SIffSS sp1 sp2) + from check_fv_cong(1)[of \1 _ _ sp1] check_fv_cong(1)[of \2 _ _ sp2] show ?thesis + unfolding SIffSS s_check_exec_simps s_check_simps fv.simps ball_conj_distrib + Iff(1)[OF 1, of sp1] Iff(3)[OF 1, of sp2] + ball_triv_nonempty[OF compatible_vals_nonemptyI[OF 1, of "fv \1 \ fv \2"]] + proof (intro arg_cong2[of _ _ _ _ "(\)"] refl, goal_cases \1 \2) + case \1 + then show ?case + by (metis (mono_tags, lifting) 1 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + next + case \2 + then show ?case + by (metis (mono_tags, lifting) 1 IntE Un_upper2 compatible_vals_extensible compatible_vals_union_eq) + qed + next + case (SIffVV vp1 vp2) + from check_fv_cong(2)[of \1 _ _ vp1] check_fv_cong(2)[of \2 _ _ vp2] show ?thesis + unfolding SIffVV s_check_exec_simps s_check_simps fv.simps ball_conj_distrib + Iff(2)[OF 1, of vp1] Iff(4)[OF 1, of vp2] + ball_triv_nonempty[OF compatible_vals_nonemptyI[OF 1, of "fv \1 \ fv \2"]] + proof (intro arg_cong2[of _ _ _ _ "(\)"] refl, goal_cases \1 \2) + case \1 + then show ?case + by (metis (mono_tags, lifting) 1 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + next + case \2 + then show ?case + by (metis (mono_tags, lifting) 1 IntE Un_upper2 compatible_vals_extensible compatible_vals_union_eq) + qed + qed (auto simp: compatible_vals_union_eq) + next + case 2 + with compatible_vals_nonemptyI[OF 2, of "fv \1 \ fv \2"] show ?case + proof (cases vp) + case (VIffSV sp1 vp2) + from check_fv_cong(1)[of \1 _ _ sp1] check_fv_cong(2)[of \2 _ _ vp2] show ?thesis + unfolding VIffSV v_check_exec_simps v_check_simps fv.simps ball_conj_distrib + Iff(1)[OF 2, of sp1] Iff(4)[OF 2, of vp2] + ball_triv_nonempty[OF compatible_vals_nonemptyI[OF 2, of "fv \1 \ fv \2"]] + proof (intro arg_cong2[of _ _ _ _ "(\)"] refl, goal_cases \1 \2) + case \1 + then show ?case + by (metis (mono_tags, lifting) 2 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + next + case \2 + then show ?case + by (metis (mono_tags, lifting) 2 IntE Un_upper2 compatible_vals_extensible compatible_vals_union_eq) + qed + next + case (VIffVS vp1 sp2) + from check_fv_cong(2)[of \1 _ _ vp1] check_fv_cong(1)[of \2 _ _ sp2] show ?thesis + unfolding VIffVS v_check_exec_simps v_check_simps fv.simps ball_conj_distrib + Iff(2)[OF 2, of vp1] Iff(3)[OF 2, of sp2] + ball_triv_nonempty[OF compatible_vals_nonemptyI[OF 2, of "fv \1 \ fv \2"]] + proof (intro arg_cong2[of _ _ _ _ "(\)"] refl, goal_cases \1 \2) + case \1 + then show ?case + by (metis (mono_tags, lifting) 2 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + next + case \2 + then show ?case + by (metis (mono_tags, lifting) 2 IntE Un_upper2 compatible_vals_extensible compatible_vals_union_eq) + qed + qed (auto simp: compatible_vals_union_eq) + } +next + case (Exists x \) + { + case 1 + then have "(vs(x := Z)) y \ {}" if "Z \ {}" for Z y + using that by auto + with 1 have IH: + "s_check_exec (vs(x := {z})) \ sp = (\v\compatible_vals (fv \) (vs(x := {z})). s_check v \ sp)" + for z sp + by (intro Exists; + auto simp: compatible_vals_fun_upd fun_upd_same + simp del: fun_upd_apply intro: fun_upd_in_compatible_vals) + from 1 show ?case + using compatible_vals_nonemptyI[OF 1, of "fv \ - {x}"] + by (cases sp) (auto simp: SubsVals_nonempty IH fun_upd_in_compatible_vals_notin compatible_vals_fun_upd) + next + case 2 + then have "(vs(x := Z)) y \ {}" if "Z \ {}" for Z y + using that by auto + with 2 have IH: + "Z \ {} \ v_check_exec (vs(x := Z)) \ vp = (\v\compatible_vals (fv \) (vs(x := Z)). v_check v \ vp)" + for Z vp + by (intro Exists; + auto simp: compatible_vals_fun_upd fun_upd_same + simp del: fun_upd_apply intro: fun_upd_in_compatible_vals) + show ?case + using compatible_vals_nonemptyI[OF 2, of "fv \ - {x}"] + by (cases vp) + (auto simp: SubsVals_nonempty IH[OF SubsVals_nonempty] + fun_upd_in_compatible_vals fun_upd_in_compatible_vals_notin compatible_vals_fun_upd + ball_conj_distrib 2[simplified] split: prod.splits if_splits | + drule bspec, assumption)+ + } +next + case (Forall x \) + { + case 1 + then have "(vs(x := Z)) y \ {}" if "Z \ {}" for Z y + using that by auto + with 1 have IH: + "Z \ {} \ s_check_exec (vs(x := Z)) \ sp = (\v\compatible_vals (fv \) (vs(x := Z)). s_check v \ sp)" + for Z sp + by (intro Forall; + auto simp: compatible_vals_fun_upd fun_upd_same + simp del: fun_upd_apply intro: fun_upd_in_compatible_vals) + show ?case + using compatible_vals_nonemptyI[OF 1, of "fv \ - {x}"] + by (cases sp) + (auto simp: SubsVals_nonempty IH[OF SubsVals_nonempty] + fun_upd_in_compatible_vals fun_upd_in_compatible_vals_notin compatible_vals_fun_upd + ball_conj_distrib 1[simplified] split: prod.splits if_splits | + drule bspec, assumption)+ + next + case 2 + then have "(vs(x := Z)) y \ {}" if "Z \ {}" for Z y + using that by auto + with 2 have IH: + "v_check_exec (vs(x := {z})) \ vp = (\v\compatible_vals (fv \) (vs(x := {z})). v_check v \ vp)" + for z vp + by (intro Forall; + auto simp: compatible_vals_fun_upd fun_upd_same + simp del: fun_upd_apply intro: fun_upd_in_compatible_vals) + from 2 show ?case + using compatible_vals_nonemptyI[OF 2, of "fv \ - {x}"] + by (cases vp) (auto simp: SubsVals_nonempty IH fun_upd_in_compatible_vals_notin compatible_vals_fun_upd) + } +next + case (Prev I \) + { + case 1 + with Prev[of vs] show ?case + using compatible_vals_nonemptyI[OF 1, of "fv \"] + by (cases sp) auto + next + case 2 + with Prev[of vs] show ?case + using compatible_vals_nonemptyI[OF 2, of "fv \"] + by (cases vp) auto + } +next + case (Next I \) + { + case 1 + with Next[of vs] show ?case + using compatible_vals_nonemptyI[OF 1, of "fv \"] + by (cases sp) (auto simp: Let_def) + next + case 2 + with Next[of vs] show ?case + using compatible_vals_nonemptyI[OF 2, of "fv \"] + by (cases vp) auto + } +next + case (Once I \) + { + case 1 + with Once[of vs] show ?case + using compatible_vals_nonemptyI[OF 1, of "fv \"] + by (cases sp) (auto simp: Let_def) + next + case 2 + with Once[of vs] show ?case + using compatible_vals_nonemptyI[OF 2, of "fv \"] + by (cases vp) auto + } +next + case (Historically I \) + { + case 1 + with Historically[of vs] show ?case + using compatible_vals_nonemptyI[OF 1, of "fv \"] + by (cases sp) auto + next + case 2 + with Historically[of vs] show ?case + using compatible_vals_nonemptyI[OF 2, of "fv \"] + by (cases vp) (auto simp: Let_def) + } +next + case (Eventually I \) + { + case 1 + with Eventually[of vs] show ?case + using compatible_vals_nonemptyI[OF 1, of "fv \"] + by (cases sp) (auto simp: Let_def) + next + case 2 + with Eventually[of vs] show ?case + using compatible_vals_nonemptyI[OF 2, of "fv \"] + by (cases vp) auto + } +next + case (Always I \) + { + case 1 + with Always[of vs] show ?case + using compatible_vals_nonemptyI[OF 1, of "fv \"] + by (cases sp) auto + next + case 2 + with Always[of vs] show ?case + using compatible_vals_nonemptyI[OF 2, of "fv \"] + by (cases vp) (auto simp: Let_def) + } +next + case (Since \1 I \2) + { + case 1 + with compatible_vals_nonemptyI[OF 1, of "fv \1 \ fv \2"] show ?case + proof (cases sp) + case (SSince sp' sps) + from check_fv_cong(1)[of \2 _ _ sp'] show ?thesis + unfolding SSince s_check_exec_simps s_check_simps fv.simps ball_conj_distrib ball_swap[of _ "set sps"] + Since(1)[OF 1] Since(3)[OF 1, of sp'] Let_def + ball_triv_nonempty[OF compatible_vals_nonemptyI[OF 1, of "fv \1 \ fv \2"]] + proof (intro arg_cong2[of _ _ _ _ "(\)"] ball_cong[of "set sps", OF refl] refl, goal_cases \2 \1) + case \2 + then show ?case + by (metis (mono_tags, lifting) 1 IntE Un_upper2 compatible_vals_extensible compatible_vals_union_eq) + next + case (\1 sp) + then show ?case using check_fv_cong(1)[of \1 _ _ sp] + by (metis (mono_tags, lifting) 1 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + qed + qed (auto simp: compatible_vals_union_eq) + next + case 2 + with compatible_vals_nonemptyI[OF 2, of "fv \1 \ fv \2"] show ?case + proof (cases vp) + case (VSince i vp' vps) + from check_fv_cong(2)[of \1 _ _ vp'] show ?thesis + unfolding VSince v_check_exec_simps v_check_simps fv.simps ball_conj_distrib ball_swap[of _ "set vps"] + Since(2)[OF 2, of vp'] Since(4)[OF 2] Let_def + ball_triv_nonempty[OF compatible_vals_nonemptyI[OF 2, of "fv \1 \ fv \2"]] + proof (intro arg_cong2[of _ _ _ _ "(\)"] ball_cong[of "set vps", OF refl] refl, goal_cases \1 \2) + case \1 + then show ?case + by (metis (mono_tags, lifting) 2 IntE Un_upper2 compatible_vals_extensible compatible_vals_union_eq) + next + case (\2 vp) + then show ?case using check_fv_cong(2)[of \2 _ _ vp] + by (metis (mono_tags, lifting) 2 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + qed + next + case (VSinceInf i j vps) + show ?thesis + unfolding VSinceInf v_check_exec_simps v_check_simps fv.simps ball_conj_distrib ball_swap[of _ "set vps"] + Since(4)[OF 2] Let_def + ball_triv_nonempty[OF compatible_vals_nonemptyI[OF 2, of "fv \1 \ fv \2"]] + proof (intro arg_cong2[of _ _ _ _ "(\)"] ball_cong[of "set vps", OF refl] refl, goal_cases \2) + case (\2 vp) + then show ?case using check_fv_cong(2)[of \2 _ _ vp] + by (metis (mono_tags, lifting) 2 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + qed + qed (auto simp: compatible_vals_union_eq) + } +next + case (Until \1 I \2) + { + case 1 + with compatible_vals_nonemptyI[OF 1, of "fv \1 \ fv \2"] show ?case + proof (cases sp) + case (SUntil sps sp') + from check_fv_cong(1)[of \2 _ _ sp'] show ?thesis + unfolding SUntil s_check_exec_simps s_check_simps fv.simps ball_conj_distrib ball_swap[of _ "set sps"] + Until(1)[OF 1] Until(3)[OF 1, of sp'] Let_def + ball_triv_nonempty[OF compatible_vals_nonemptyI[OF 1, of "fv \1 \ fv \2"]] + proof (intro arg_cong2[of _ _ _ _ "(\)"] ball_cong[of "set sps", OF refl] refl, goal_cases \2 \1) + case \2 + then show ?case + by (metis (mono_tags, lifting) 1 IntE Un_upper2 compatible_vals_extensible compatible_vals_union_eq) + next + case (\1 sp) + then show ?case using check_fv_cong(1)[of \1 _ _ sp] + by (metis (mono_tags, lifting) 1 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + qed + qed (auto simp: compatible_vals_union_eq) + next + case 2 + with compatible_vals_nonemptyI[OF 2, of "fv \1 \ fv \2"] show ?case + proof (cases vp) + case (VUntil i vps vp') + from check_fv_cong(2)[of \1 _ _ vp'] show ?thesis + unfolding VUntil v_check_exec_simps v_check_simps fv.simps ball_conj_distrib ball_swap[of _ "set vps"] + Until(2)[OF 2, of vp'] Until(4)[OF 2] Let_def + ball_triv_nonempty[OF compatible_vals_nonemptyI[OF 2, of "fv \1 \ fv \2"]] + proof (intro arg_cong2[of _ _ _ _ "(\)"] ball_cong[of "set vps", OF refl] refl, goal_cases \1 \2) + case \1 + then show ?case + by (metis (mono_tags, lifting) 2 IntE Un_upper2 compatible_vals_extensible compatible_vals_union_eq) + next + case (\2 vp) + then show ?case using check_fv_cong(2)[of \2 _ _ vp] + by (metis (mono_tags, lifting) 2 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + qed + next + case (VUntilInf i j vps) + show ?thesis + unfolding VUntilInf v_check_exec_simps v_check_simps fv.simps ball_conj_distrib ball_swap[of _ "set vps"] + Until(4)[OF 2] Let_def + ball_triv_nonempty[OF compatible_vals_nonemptyI[OF 2, of "fv \1 \ fv \2"]] + proof (intro arg_cong2[of _ _ _ _ "(\)"] ball_cong[of "set vps", OF refl] refl, goal_cases \2) + case (\2 vp) + then show ?case using check_fv_cong(2)[of \2 _ _ vp] + by (metis (mono_tags, lifting) 2 IntE Un_upper1 compatible_vals_extensible compatible_vals_union_eq) + qed + qed (auto simp: compatible_vals_union_eq) + } +qed + +lemma s_check_code[code]: "s_check v \ sp = s_check_exec (\x. {v x}) \ sp" + by (subst check_exec_check) + (auto simp: compatible_vals_def elim: check_fv_cong[THEN iffD2, rotated]) + +lemma v_check_code[code]: "v_check v \ vp = v_check_exec (\x. {v x}) \ vp" + by (subst check_exec_check) + (auto simp: compatible_vals_def elim: check_fv_cong[THEN iffD2, rotated]) + +subsection \Latest Relevant Time-Point\ + +fun LRTP :: "('n, 'd) formula \ nat \ nat option" where + "LRTP \ i = Some i" +| "LRTP \ i = Some i" +| "LRTP (_ \ _) i = Some i" +| "LRTP (_ \<^bold>\ _) i = Some i" +| "LRTP (\\<^sub>F \) i = LRTP \ i" +| "LRTP (\ \\<^sub>F \) i = max_opt (LRTP \ i) (LRTP \ i)" +| "LRTP (\ \\<^sub>F \) i = max_opt (LRTP \ i) (LRTP \ i)" +| "LRTP (\ \\<^sub>F \) i = max_opt (LRTP \ i) (LRTP \ i)" +| "LRTP (\ \\<^sub>F \) i = max_opt (LRTP \ i) (LRTP \ i)" +| "LRTP (\\<^sub>F_. \) i = LRTP \ i" +| "LRTP (\\<^sub>F_. \) i = LRTP \ i" +| "LRTP (\<^bold>Y I \) i = LRTP \ (i-1)" +| "LRTP (\<^bold>X I \) i = LRTP \ (i+1)" +| "LRTP (\<^bold>P I \) i = LRTP \ (LTP_p_safe \ i I)" +| "LRTP (\<^bold>H I \) i = LRTP \ (LTP_p_safe \ i I)" +| "LRTP (\<^bold>F I \) i = (case right I of \ \ None | enat b \ LRTP \ (LTP_f \ i b))" +| "LRTP (\<^bold>G I \) i = (case right I of \ \ None | enat b \ LRTP \ (LTP_f \ i b))" +| "LRTP (\ \<^bold>S I \) i = max_opt (LRTP \ i) (LRTP \ (LTP_p_safe \ i I))" +| "LRTP (\ \<^bold>U I \) i = (case right I of \ \ None | enat b \ max_opt (LRTP \ ((LTP_f \ i b)-1)) (LRTP \ (LTP_f \ i b)))" + +lemma fb_LRTP: + assumes "future_bounded \" + shows "\ Option.is_none (LRTP \ i)" + using assms + by (induction \ i rule: LRTP.induct) + (auto simp add: max_opt_def Option.is_none_def) + +lemma not_none_fb_LRTP: + assumes "future_bounded \" + shows "LRTP \ i \ None" + using assms fb_LRTP by (auto simp add: Option.is_none_def) + +lemma is_some_fb_LRTP: + assumes "future_bounded \" + shows "\j. LRTP \ i = Some j" + using assms fb_LRTP by (auto simp add: Option.is_none_def) + +lemma enat_trans[simp]: "enat i \ enat j \ enat j \ enat k \ enat i \ enat k" + by auto + +subsection \Active Domain\ + +definition AD :: "('n, 'd) formula \ nat \ 'd set" + where "AD \ i = consts \ \ (\ k \ the (LRTP \ i). \ (set ` snd ` \ \ k))" + +lemma val_in_AD_iff: + "x \ fv \ \ v x \ AD \ i \ v x \ consts \ \ + (\r ts k. k \ the (LRTP \ i) \ (r, v\<^bold>\ts\<^bold>\) \ \ \ k \ x \ \ (set (map fv_trm ts)))" + unfolding AD_def Un_iff UN_iff Bex_def atMost_iff set_map + ex_comm[of "P :: _ \ nat \ _" for P] ex_simps image_iff +proof (safe intro!: arg_cong[of _ _ "\x. _ \ x"] ex_cong, unfold snd_conv, goal_cases LR RL) + case (LR i _ r ds) + then show ?case + by (intro exI[of _ r] conjI + exI[of _ "map (\d. if v x = d then (\<^bold>v x) else \<^bold>c d) ds"]) + (auto simp: eval_trms_def o_def map_idI) +next + case (RL i r ts t) + then show ?case + by (intro exI[of _ "v\<^bold>\ts\<^bold>\"] conjI) + (auto simp: eval_trms_def image_iff in_fv_trm_conv intro!: bexI[of _ t]) +qed + +lemma val_notin_AD_iff: + "x \ fv \ \ v x \ AD \ i \ v x \ consts \ \ + (\r ts k. k \ the (LRTP \ i) \ x \ \ (set (map fv_trm ts)) \ (r, v\<^bold>\ts\<^bold>\) \ \ \ k)" + using val_in_AD_iff by blast + +lemma finite_values: "finite (\ (set ` snd ` \ \ k))" + by (transfer, auto simp add: sfinite_def) + +lemma finite_tps: "future_bounded \ \ finite (\ k < the (LRTP \ i). {k})" + using fb_LRTP[of \] finite_enat_bounded + by simp + +lemma finite_AD [simp]: "future_bounded \ \ finite (AD \ i)" + using finite_tps finite_values + by (simp add: AD_def enat_def) + +lemma finite_AD_UNIV: + assumes "future_bounded \" and "AD \ i = (UNIV:: 'd set)" + shows "finite (UNIV::'d set)" +proof - + have "finite (AD \ i)" + using finite_AD[of \ i, OF assms(1)] by simp + then show ?thesis + using assms(2) by simp +qed + +subsection \Congruence Modulo Active Domain\ + +lemma AD_simps[simp]: + "AD (\\<^sub>F \) i = AD \ i" + "future_bounded (\ \\<^sub>F \) \ AD (\ \\<^sub>F \) i = AD \ i \ AD \ i" + "future_bounded (\ \\<^sub>F \) \ AD (\ \\<^sub>F \) i = AD \ i \ AD \ i" + "future_bounded (\ \\<^sub>F \) \ AD (\ \\<^sub>F \) i = AD \ i \ AD \ i" + "future_bounded (\ \\<^sub>F \) \ AD (\ \\<^sub>F \) i = AD \ i \ AD \ i" + "AD (\\<^sub>Fx. \) i = AD \ i" + "AD (\\<^sub>Fx. \) i = AD \ i" + "AD (\<^bold>Y I \) i = AD \ (i - 1)" + "AD (\<^bold>X I \) i = AD \ (i + 1)" + "future_bounded (\<^bold>F I \) \ AD (\<^bold>F I \) i = AD \ (LTP_f \ i (the_enat (right I)))" + "future_bounded (\<^bold>G I \) \ AD (\<^bold>G I \) i = AD \ (LTP_f \ i (the_enat (right I)))" + "AD (\<^bold>P I \) i = AD \ (LTP_p_safe \ i I)" + "AD (\<^bold>H I \) i = AD \ (LTP_p_safe \ i I)" + "future_bounded (\ \<^bold>S I \) \ AD (\ \<^bold>S I \) i = AD \ i \ AD \ (LTP_p_safe \ i I)" + "future_bounded (\ \<^bold>U I \) \ AD (\ \<^bold>U I \) i = AD \ (LTP_f \ i (the_enat (right I)) - 1) \ AD \ (LTP_f \ i (the_enat (right I)))" + by (auto 0 3 simp: AD_def max_opt_def not_none_fb_LRTP le_max_iff_disj Bex_def split: option.splits) + + +lemma LTP_p_mono: "i \ j \ LTP_p_safe \ i I \ LTP_p_safe \ j I" + unfolding LTP_p_safe_def + by (smt (verit, ccfv_threshold) \_mono bot_nat_0.extremum diff_le_mono order.trans i_LTP_tau le_cases3 min.bounded_iff) + +lemma LTP_f_mono: + assumes "i \ j" + shows"LTP_f \ i b \ LTP_f \ j b" + unfolding LTP_def +proof (rule Max_mono) + show "finite {i. \ \ i \ \ \ j + b}" + unfolding finite_nat_set_iff_bounded_le + by (metis i_le_LTPi_add le_Suc_ex mem_Collect_eq) +qed (auto simp: assms intro!: exI[of _ i] elim: order_trans) + +lemma LRTP_mono: "future_bounded \ \ i \ j \ the (LRTP \ i) \ the (LRTP \ j)" +proof (induct \ arbitrary: i j) + case (Or \1 \2) + from Or(1,2)[of i j] Or(3-) show ?case + by (auto simp: max_opt_def not_none_fb_LRTP split: option.splits) +next + case (And \1 \2) + from And(1,2)[of i j] And(3-) show ?case + by (auto simp: max_opt_def not_none_fb_LRTP split: option.splits) +next + case (Imp \1 \2) + from Imp(1,2)[of i j] Imp(3-) show ?case + by (auto simp: max_opt_def not_none_fb_LRTP split: option.splits) +next + case (Iff \1 \2) + from Iff(1,2)[of i j] Iff(3-) show ?case + by (auto simp: max_opt_def not_none_fb_LRTP split: option.splits) +next + case (Since \1 I \2) + from Since(1)[OF _ Since(4)] Since(2)[of "LTP_p_safe \ i I" "LTP_p_safe \ j I"] Since(3-) + show ?case + by (auto simp: max_opt_def not_none_fb_LRTP LTP_p_mono split: option.splits) +next + case (Until \1 I \2) + from Until(1)[of "LTP_f \ i (the_enat (right I)) - 1" "LTP_f \ j (the_enat (right I)) - 1"] + Until(2)[of "LTP_f \ i (the_enat (right I))" "LTP_f \ j (the_enat (right I))"] Until(3-) + show ?case + by (auto simp: max_opt_def not_none_fb_LRTP LTP_f_mono diff_le_mono split: option.splits) +qed (auto simp: LTP_p_mono LTP_f_mono) + +lemma AD_mono: "future_bounded \ \ i \ j \ AD \ i \ AD \ j" + by (auto 0 3 simp: AD_def Bex_def intro: LRTP_mono elim!: order_trans) + +lemma LTP_p_safe_le[simp]: "LTP_p_safe \ i I \ i" + by (auto simp: LTP_p_safe_def) + +lemma check_AD_cong: + assumes "future_bounded \" + and "(\x \ fv \. v x = v' x \ (v x \ AD \ i \ v' x \ AD \ i))" + shows "(s_at sp = i \ s_check v \ sp \ s_check v' \ sp)" + "(v_at vp = i \ v_check v \ vp \ v_check v' \ vp)" + using assms +proof (induction v \ sp and v \ vp arbitrary: i v' and i v' rule: s_check_v_check.induct) + case (1 v f sp) + note IH = 1(1-23)[OF refl] and hyps = 1(24-26) + show ?case + proof (cases sp) + case (SPred j r ts) + then show ?thesis + proof (cases f) + case (Pred q us) + with SPred hyps show ?thesis + using eval_trms_fv_cong[of ts v v'] + by (force simp: val_notin_AD_iff dest!: spec[of _ i] spec[of _ r] spec[of _ ts]) + qed auto + next + case (SEq_Const j r ts) + with hyps show ?thesis + by (cases f) (auto simp: val_notin_AD_iff) + next + case (SNeg vp') + then show ?thesis + using IH(1)[of _ _ _ v'] hyps + by (cases f) auto + next + case (SOrL sp') + then show ?thesis + using IH(2)[of _ _ _ _ v'] hyps + by (cases f) auto + next + case (SOrR sp') + then show ?thesis + using IH(3)[of _ _ _ _ v'] hyps + by (cases f) auto + next + case (SAnd sp1 sp2) + then show ?thesis + using IH(4,5)[of _ _ _ _ _ v'] hyps + by (cases f) (auto 7 0)+ + next + case (SImpL vp') + then show ?thesis + using IH(6)[of _ _ _ _ v'] hyps + by (cases f) auto + next + case (SImpR sp') + then show ?thesis + using IH(7)[of _ _ _ _ v'] hyps + by (cases f) auto + next + case (SIffSS sp1 sp2) + then show ?thesis + using IH(8,9)[of _ _ _ _ _ v'] hyps + by (cases f) (auto 7 0)+ + next + case (SIffVV vp1 vp2) + then show ?thesis + using IH(10,11)[of _ _ _ _ _ v'] hyps + by (cases f) (auto 7 0)+ + next + case (SExists x z sp') + then show ?thesis + using IH(12)[of x _ x z sp' i "v'(x := z)"] hyps + by (cases f) (auto simp add: fun_upd_def) + next + case (SForall x part) + then show ?thesis + using IH(13)[of x _ x part _ _ D _ z _ "v'(x := z)" for D z, OF _ _ _ _ refl _ refl] hyps + by (cases f) (auto simp add: fun_upd_def) + next + case (SPrev sp') + then show ?thesis + using IH(14)[of _ _ _ _ _ _ v'] hyps + by (cases f) auto + next + case (SNext sp') + then show ?thesis + using IH(15)[of _ _ _ _ _ _ v'] hyps + by (cases f) (auto simp add: Let_def) + next + case (SOnce j sp') + then show ?thesis + proof (cases f) + case (Once I \) + { fix k + assume k: "k \ i" "\ \ i - left I \ \ \ k" + then have "\ \ i - left I \ \ \ 0" + by (meson \_mono le0 order_trans) + with k have "k \ LTP_p_safe \ i I" + unfolding LTP_p_safe_def by (auto simp: i_LTP_tau) + with Once hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ k \ v' x \ AD \ k" + by (auto dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + with Once SOnce show ?thesis + using IH(16)[OF Once SOnce refl refl, of v'] hyps(1,2) + by (auto simp: Let_def le_diff_conv2) + qed auto + next + case (SHistorically j k sps) + then show ?thesis + proof (cases f) + case (Historically I \) + { fix sp :: "('n, 'd) sproof" + define l and u where "l = s_at sp" and "u = LTP_p \ i I" + assume *: "sp \ set sps" "\ \ 0 + left I \ \ \ i" + then have u_def: "u = LTP_p_safe \ i I" + by (auto simp: LTP_p_safe_def u_def) + from *(1) obtain j where j: "sp = sps ! j" "j < length sps" + unfolding in_set_conv_nth by auto + moreover + assume eq: "map s_at sps = [k ..< Suc u]" + then have len: "length sps = Suc u - k" + by (auto dest!: arg_cong[where f=length]) + moreover + have "s_at (sps ! j) = k + j" + using arg_cong[where f="\xs. nth xs j", OF eq] j len *(2) + by (auto simp: nth_append) + ultimately have "l \ u" + unfolding l_def by auto + with Historically hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ l \ v' x \ AD \ l" + by (auto simp: u_def dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + with Historically SHistorically show ?thesis + using IH(17)[OF Historically SHistorically _ refl, of _ v'] hyps(1,2) + by auto + qed auto + next + case (SEventually j sp') + then show ?thesis + proof (cases f) + case (Eventually I \) + { fix k + assume "\ \ k \ the_enat (right I) + \ \ i" + then have "k \ LTP_f \ i (the_enat (right I))" + by (metis add.commute i_le_LTPi_add le_add_diff_inverse) + with Eventually hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ k \ v' x \ AD \ k" + by (auto dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + with Eventually SEventually show ?thesis + using IH(18)[OF Eventually SEventually refl refl, of v'] hyps(1,2) + by (auto simp: Let_def) + qed auto + next + case (SAlways j k sps) + then show ?thesis + proof (cases f) + case (Always I \) + { fix sp :: "('n, 'd) sproof" + define l and u where "l = s_at sp" and "u = LTP_f \ i (the_enat (right I))" + assume *: "sp \ set sps" + then obtain j where j: "sp = sps ! j" "j < length sps" + unfolding in_set_conv_nth by auto + assume eq: "map s_at sps = [ETP_f \ i I ..< Suc u]" + then have "length sps = Suc u - ETP_f \ i I" + by (auto dest!: arg_cong[where f=length]) + with j eq have "l \ LTP_f \ i (the_enat (right I))" + by (auto simp: l_def u_def dest!: arg_cong[where f="\xs. nth xs j"] + simp del: upt.simps split: if_splits) + with Always hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ l \ v' x \ AD \ l" + by (auto dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + with Always SAlways show ?thesis + using IH(19)[OF Always SAlways _ refl, of _ v'] hyps(1,2) + by auto + qed auto + next + case (SSince sp' sps) + then show ?thesis + proof (cases f) + case (Since \ I \) + { fix sp :: "('n, 'd) sproof" + define l where "l = s_at sp" + assume *: "sp \ set sps" + from *(1) obtain j where j: "sp = sps ! j" "j < length sps" + unfolding in_set_conv_nth by auto + moreover + assume eq: "map s_at sps = [Suc (s_at sp') ..< Suc i]" + then have len: "length sps = i - s_at sp'" + by (auto dest!: arg_cong[where f=length]) + moreover + have "s_at (sps ! j) = Suc (s_at sp') + j" + using arg_cong[where f="\xs. nth xs j", OF eq] j len + by (auto simp: nth_append) + ultimately have "l \ i" + unfolding l_def by auto + with Since hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ l \ v' x \ AD \ l" + by (auto simp: dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + moreover + { fix k + assume k: "k \ i" "\ \ i - left I \ \ \ k" + then have "\ \ i - left I \ \ \ 0" + by (meson \_mono le0 order_trans) + with k have "k \ LTP_p_safe \ i I" + unfolding LTP_p_safe_def by (auto simp: i_LTP_tau) + with Since hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ k \ v' x \ AD \ k" + by (auto dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + ultimately show ?thesis + using Since SSince IH(20)[OF Since SSince refl refl refl, of v'] IH(21)[OF Since SSince refl refl _ refl, of _ v'] hyps(1,2) + by (auto simp: Let_def le_diff_conv2 simp del: upt.simps) + qed auto + next + case (SUntil sps sp') + then show ?thesis + proof (cases f) + case (Until \ I \) + { fix sp :: "('n, 'd) sproof" + define l where "l = s_at sp" + assume *: "sp \ set sps" + from *(1) obtain j where j: "sp = sps ! j" "j < length sps" + unfolding in_set_conv_nth by auto + moreover + assume "\ \ (s_at sp') i \ the_enat (right I)" + then have "s_at sp' \ LTP_f \ i (the_enat (right I))" + by (metis add.commute i_le_LTPi_add le_add_diff_inverse le_diff_conv) + moreover + assume eq: "map s_at sps = [i ..< s_at sp']" + then have len: "length sps = s_at sp' - i" + by (auto dest!: arg_cong[where f=length]) + moreover + have "s_at (sps ! j) = i + j" + using arg_cong[where f="\xs. nth xs j", OF eq] j len + by (auto simp: nth_append) + ultimately have "l \ LTP_f \ i (the_enat (right I)) - 1" + unfolding l_def by auto + with Until hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ l \ v' x \ AD \ l" + by (auto simp: dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + moreover + { fix k + assume "\ \ k \ the_enat (right I) + \ \ i" + then have "k \ LTP_f \ i (the_enat (right I))" + by (metis add.commute i_le_LTPi_add le_add_diff_inverse) + with Until hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ k \ v' x \ AD \ k" + by (auto dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + ultimately show ?thesis + using Until SUntil IH(22)[OF Until SUntil refl refl refl, of v'] IH(23)[OF Until SUntil refl refl _ refl, of _ v'] hyps(1,2) + by (auto simp: Let_def le_diff_conv2 simp del: upt.simps) + qed auto + qed (cases f; simp_all)+ +next + case (2 v f vp) + note IH = 2(1-25)[OF refl] and hyps = 2(26-28) + show ?case + proof (cases vp) + case (VPred j r ts) + then show ?thesis + proof (cases f) + case (Pred q us) + with VPred hyps show ?thesis + using eval_trms_fv_cong[of ts v v'] + by (force simp: val_notin_AD_iff dest!: spec[of _ i] spec[of _ r] spec[of _ ts]) + qed auto + next + case (VEq_Const j r ts) + with hyps show ?thesis + by (cases f) (auto simp: val_notin_AD_iff) + next + case (VNeg sp') + then show ?thesis + using IH(1)[of _ _ _ v'] hyps + by (cases f) auto + next + case (VOr vp1 vp2) + then show ?thesis + using IH(2,3)[of _ _ _ _ _ v'] hyps + by (cases f) (auto 7 0)+ + next + case (VAndL vp') + then show ?thesis + using IH(4)[of _ _ _ _ v'] hyps + by (cases f) auto + next + case (VAndR vp') + then show ?thesis + using IH(5)[of _ _ _ _ v'] hyps + by (cases f) auto + next + case (VImp sp1 vp2) + then show ?thesis + using IH(6,7)[of _ _ _ _ _ v'] hyps + by (cases f) (auto 7 0)+ + next + case (VIffSV sp1 vp2) + then show ?thesis + using IH(8,9)[of _ _ _ _ _ v'] hyps + by (cases f) (auto 7 0)+ + next + case (VIffVS vp1 sp2) + then show ?thesis + using IH(10,11)[of _ _ _ _ _ v'] hyps + by (cases f) (auto 7 0)+ + next + case (VExists x part) + then show ?thesis + using IH(12)[of x _ x part _ _ D _ z _ "v'(x := z)" for D z, OF _ _ _ _ refl _ refl] hyps + by (cases f) (auto simp add: fun_upd_def) + next + case (VForall x z vp') + then show ?thesis + using IH(13)[of x _ x z vp' i "v'(x := z)"] hyps + by (cases f) (auto simp add: fun_upd_def) + next + case (VPrev vp') + then show ?thesis + using IH(14)[of _ _ _ _ _ _ v'] hyps + by (cases f) auto + next + case (VNext vp') + then show ?thesis + using IH(15)[of _ _ _ _ _ _ v'] hyps + by (cases f) auto + next + case (VOnce j k vps) + then show ?thesis + proof (cases f) + case (Once I \) + { fix vp :: "('n, 'd) vproof" + define l and u where "l = v_at vp" and "u = LTP_p \ i I" + assume *: "vp \ set vps" "\ \ 0 + left I \ \ \ i" + then have u_def: "u = LTP_p_safe \ i I" + by (auto simp: LTP_p_safe_def u_def) + from *(1) obtain j where j: "vp = vps ! j" "j < length vps" + unfolding in_set_conv_nth by auto + moreover + assume eq: "map v_at vps = [k ..< Suc u]" + then have len: "length vps = Suc u - k" + by (auto dest!: arg_cong[where f=length]) + moreover + have "v_at (vps ! j) = k + j" + using arg_cong[where f="\xs. nth xs j", OF eq] j len *(2) + by (auto simp: nth_append) + ultimately have "l \ u" + unfolding l_def by auto + with Once hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ l \ v' x \ AD \ l" + by (auto simp: u_def dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + with Once VOnce show ?thesis + using IH(16)[OF Once VOnce _ refl, of _ v'] hyps(1,2) + by auto + qed auto + next + case (VHistorically j vp') + then show ?thesis + proof (cases f) + case (Historically I \) + { fix k + assume k: "k \ i" "\ \ i - left I \ \ \ k" + then have "\ \ i - left I \ \ \ 0" + by (meson \_mono le0 order_trans) + with k have "k \ LTP_p_safe \ i I" + unfolding LTP_p_safe_def by (auto simp: i_LTP_tau) + with Historically hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ k \ v' x \ AD \ k" + by (auto dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + with Historically VHistorically show ?thesis + using IH(17)[OF Historically VHistorically refl refl, of v'] hyps(1,2) + by (auto simp: Let_def le_diff_conv2) + qed auto + next + case (VEventually j k vps) + then show ?thesis + proof (cases f) + case (Eventually I \) + { fix vp :: "('n, 'd) vproof" + define l and u where "l = v_at vp" and "u = LTP_f \ i (the_enat (right I))" + assume *: "vp \ set vps" + then obtain j where j: "vp = vps ! j" "j < length vps" + unfolding in_set_conv_nth by auto + assume eq: "map v_at vps = [ETP_f \ i I ..< Suc u]" + then have "length vps = Suc u - ETP_f \ i I" + by (auto dest!: arg_cong[where f=length]) + with j eq have "l \ LTP_f \ i (the_enat (right I))" + by (auto simp: l_def u_def dest!: arg_cong[where f="\xs. nth xs j"] + simp del: upt.simps split: if_splits) + with Eventually hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ l \ v' x \ AD \ l" + by (auto dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + with Eventually VEventually show ?thesis + using IH(18)[OF Eventually VEventually _ refl, of _ v'] hyps(1,2) + by auto + qed auto + next + case (VAlways j vp') + then show ?thesis + proof (cases f) + case (Always I \) + { fix k + assume "\ \ k \ the_enat (right I) + \ \ i" + then have "k \ LTP_f \ i (the_enat (right I))" + by (metis add.commute i_le_LTPi_add le_add_diff_inverse) + with Always hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ k \ v' x \ AD \ k" + by (auto dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + with Always VAlways show ?thesis + using IH(19)[OF Always VAlways refl refl, of v'] hyps(1,2) + by (auto simp: Let_def) + qed auto + next + case (VSince j vp' vps) + then show ?thesis + proof (cases f) + case (Since \ I \) + { fix sp :: "('n, 'd) vproof" + define l and u where "l = v_at sp" and "u = LTP_p \ i I" + assume *: "sp \ set vps" "\ \ 0 + left I \ \ \ i" + then have u_def: "u = LTP_p_safe \ i I" + by (auto simp: LTP_p_safe_def u_def) + from *(1) obtain j where j: "sp = vps ! j" "j < length vps" + unfolding in_set_conv_nth by auto + moreover + assume eq: "map v_at vps = [v_at vp' ..< Suc u]" + then have len: "length vps = Suc u - v_at vp'" + by (auto dest!: arg_cong[where f=length]) + moreover + have "v_at (vps ! j) = v_at vp' + j" + using arg_cong[where f="\xs. nth xs j", OF eq] j len + by (auto simp: nth_append) + ultimately have "l \ u" + unfolding l_def by auto + with Since hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ l \ v' x \ AD \ l" + by (auto simp: u_def dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + moreover + { fix k + assume k: "k \ i" + with Since hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ k \ v' x \ AD \ k" + by (auto dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + ultimately show ?thesis + using Since VSince IH(20)[OF Since VSince refl refl, of v'] IH(21)[OF Since VSince refl _ refl, of _ v'] hyps(1,2) + by (auto simp: Let_def le_diff_conv2 simp del: upt.simps) + qed auto + next + case (VSinceInf j k vps) + then show ?thesis + proof (cases f) + case (Since \ I \) + { fix vp :: "('n, 'd) vproof" + define l and u where "l = v_at vp" and "u = LTP_p \ i I" + assume *: "vp \ set vps" "\ \ 0 + left I \ \ \ i" + then have u_def: "u = LTP_p_safe \ i I" + by (auto simp: LTP_p_safe_def u_def) + from *(1) obtain j where j: "vp = vps ! j" "j < length vps" + unfolding in_set_conv_nth by auto + moreover + assume eq: "map v_at vps = [k ..< Suc u]" + then have len: "length vps = Suc u - k" + by (auto dest!: arg_cong[where f=length]) + moreover + have "v_at (vps ! j) = k + j" + using arg_cong[where f="\xs. nth xs j", OF eq] j len *(2) + by (auto simp: nth_append) + ultimately have "l \ u" + unfolding l_def by auto + with Since hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ l \ v' x \ AD \ l" + by (auto simp: u_def dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + with Since VSinceInf show ?thesis + using IH(22)[OF Since VSinceInf _ refl, of _ v'] hyps(1,2) + by auto + qed auto + next + case (VUntil j vps vp') + then show ?thesis + proof (cases f) + case (Until \ I \) + { fix sp :: "('n, 'd) vproof" + define l and u where "l = v_at sp" and "u = v_at vp'" + assume *: "sp \ set vps" "v_at vp' \ LTP_f \ i (the_enat (right I))" + from *(1) obtain j where j: "sp = vps ! j" "j < length vps" + unfolding in_set_conv_nth by auto + moreover + assume eq: "map v_at vps = [ETP_f \ i I ..< Suc u]" + then have "length vps = Suc u - ETP_f \ i I" + by (auto dest!: arg_cong[where f=length]) + with j eq *(2) have "l \ LTP_f \ i (the_enat (right I))" + by (auto simp: l_def u_def dest!: arg_cong[where f="\xs. nth xs j"] + simp del: upt.simps split: if_splits) + with Until hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ l \ v' x \ AD \ l" + by (auto dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + moreover + { fix k + assume "k < LTP_f \ i (the_enat (right I))" + then have "k \ LTP_f \ i (the_enat (right I)) - 1" + by linarith + with Until hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ k \ v' x \ AD \ k" + by (auto dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + ultimately show ?thesis + using Until VUntil IH(23)[OF Until VUntil refl refl, of v'] IH(24)[OF Until VUntil refl _ refl, of _ v'] hyps(1,2) + by (auto simp: Let_def le_diff_conv2 simp del: upt.simps) + qed auto + next + case (VUntilInf j k vps) + then show ?thesis + proof (cases f) + case (Until \ I \) + { fix vp :: "('n, 'd) vproof" + define l and u where "l = v_at vp" and "u = LTP_f \ i (the_enat (right I))" + assume *: "vp \ set vps" + then obtain j where j: "vp = vps ! j" "j < length vps" + unfolding in_set_conv_nth by auto + assume eq: "map v_at vps = [ETP_f \ i I ..< Suc u]" + then have "length vps = Suc u - ETP_f \ i I" + by (auto dest!: arg_cong[where f=length]) + with j eq have "l \ LTP_f \ i (the_enat (right I))" + by (auto simp: l_def u_def dest!: arg_cong[where f="\xs. nth xs j"] + simp del: upt.simps split: if_splits) + with Until hyps(2,3) have "\x\fv \. v x = v' x \ v x \ AD \ l \ v' x \ AD \ l" + by (auto dest!: bspec dest: AD_mono[THEN set_mp, rotated -1]) + } + with Until VUntilInf show ?thesis + using IH(25)[OF Until VUntilInf _ refl, of _ v'] hyps(1,2) + by auto + qed auto + qed (cases f; simp_all)+ +qed + +subsection \Checker Completeness\ + +lemma part_hd_tabulate: "distinct xs \ part_hd (tabulate xs f z) = (case xs of [] \ z | (x # _) \ (if set xs = UNIV then f x else z))" + by (transfer, auto split: list.splits) + +lemma s_at_tabulate: + assumes "\z. s_at (mypick z) = i" + and "mypart = tabulate (sorted_list_of_set (AD \ i)) mypick (mypick (SOME z. z \ AD \ i))" + shows "\(sub, vp) \ SubsVals mypart. s_at vp = i" + using assms by (transfer, auto) + +lemma v_at_tabulate: + assumes "\z. v_at (mypick z) = i" + and "mypart = tabulate (sorted_list_of_set (AD \ i)) mypick (mypick (SOME z. z \ AD \ i))" + shows "\(sub, vp) \ SubsVals mypart. v_at vp = i" + using assms by (transfer, auto) + +lemma s_check_tabulate: + assumes "future_bounded \" + and "\z. s_at (mypick z) = i" + and "\z. s_check (v(x:=z)) \ (mypick z)" + and "mypart = tabulate (sorted_list_of_set (AD \ i)) mypick (mypick (SOME z. z \ AD \ i))" + shows "\(sub, vp) \ SubsVals mypart. \z \ sub. s_check (v(x := z)) \ vp" + using assms +proof (transfer fixing: \ \ mypick i v x, goal_cases 1) + case (1 mypart) + { fix z + assume s_at_assm: "\z. s_at (mypick z) = i" + and s_check_assm: "\z. s_check (v(x := z)) \ (mypick z)" + and fb_assm: "future_bounded \" + and z_notin_AD: "z \ (AD \ i)" + have s_at_mypick: "s_at (mypick (SOME z. z \ local.AD \ i)) = i" + using s_at_assm by simp + have s_check_mypick: "Checker.s_check \ (v(x := SOME z. z \ AD \ i)) \ (mypick (SOME z. z \ AD \ i))" + using s_check_assm by simp + have "s_check (v(x := z)) \ (mypick (SOME z. z \ AD \ i))" + using z_notin_AD + by (subst check_AD_cong(1)[of \ "v(x := z)" "v(x := (SOME z. z \ Checker.AD \ \ i))" i "mypick (SOME z. z \ AD \ i)", OF fb_assm _ s_at_mypick]) + (auto simp add: someI[of "\z. z \ AD \ i" z] s_check_mypick fb_assm split: if_splits) + } + with 1 show ?case + by auto +qed + +lemma v_check_tabulate: + assumes "future_bounded \" + and "\z. v_at (mypick z) = i" + and "\z. v_check (v(x:=z)) \ (mypick z)" + and "mypart = tabulate (sorted_list_of_set (AD \ i)) mypick (mypick (SOME z. z \ AD \ i))" + shows "\(sub, vp) \ SubsVals mypart. \z \ sub. v_check (v(x := z)) \ vp" + using assms +proof (transfer fixing: \ \ mypick i v x, goal_cases 1) + case (1 mypart) + { fix z + assume v_at_assm: "\z. v_at (mypick z) = i" + and v_check_assm: "\z. v_check (v(x := z)) \ (mypick z)" + and fb_assm: "future_bounded \" + and z_notin_AD: "z \ (AD \ i)" + have v_at_mypick: "v_at (mypick (SOME z. z \ local.AD \ i)) = i" + using v_at_assm by simp + have v_check_mypick: "Checker.v_check \ (v(x := SOME z. z \ AD \ i)) \ (mypick (SOME z. z \ AD \ i))" + using v_check_assm by simp + have "v_check (v(x := z)) \ (mypick (SOME z. z \ AD \ i))" + using z_notin_AD + by (subst check_AD_cong(2)[of \ "v(x := z)" "v(x := (SOME z. z \ Checker.AD \ \ i))" i "mypick (SOME z. z \ AD \ i)", OF fb_assm _ v_at_mypick]) + (auto simp add: someI[of "\z. z \ AD \ i" z] v_check_mypick fb_assm split: if_splits) + } + with 1 show ?case + by auto +qed + +lemma s_at_part_hd_tabulate: + assumes "future_bounded \" + and "\z. s_at (f z) = i" + and "mypart = tabulate (sorted_list_of_set (AD \ i)) f (f (SOME z. z \ AD \ i))" + shows "s_at (part_hd mypart) = i" + using assms by (simp add: part_hd_tabulate split: list.splits) + +lemma v_at_part_hd_tabulate: + assumes "future_bounded \" + and "\z. v_at (f z) = i" + and "mypart = tabulate (sorted_list_of_set (AD \ i)) f (f (SOME z. z \ AD \ i))" + shows "v_at (part_hd mypart) = i" + using assms by (simp add: part_hd_tabulate split: list.splits) + +lemma check_completeness_aux: + "(SAT \ v i \ \ future_bounded \ \ (\sp. s_at sp = i \ s_check v \ sp)) \ + (VIO \ v i \ \ future_bounded \ \ (\vp. v_at vp = i \ v_check v \ vp))" +proof (induct v i \ rule: SAT_VIO.induct) + case (STT v i) + then show ?case + by (auto intro!: exI[of _ "STT i"]) +next + case (VFF v i) + then show ?case + by (auto intro!: exI[of _ "VFF i"]) +next + case (SPred r v ts i) + then show ?case + by (auto intro!: exI[of _ "SPred i r ts"]) +next + case (VPred r v ts i) + then show ?case + by (auto intro!: exI[of _ "VPred i r ts"]) +next + case (SEq_Const v x c i) + then show ?case + by (auto intro!: exI[of _ "SEq_Const i x c"]) +next + case (VEq_Const v x c i) + then show ?case + by (auto intro!: exI[of _ "VEq_Const i x c"]) +next + case (SNeg v i \) + then show ?case + by (auto intro: exI[of _ "SNeg _"]) +next + case (VNeg v i \) + then show ?case + by (auto intro: exI[of _ "VNeg _"]) +next + case (SOrL v i \ \) + then show ?case + by (auto intro: exI[of _ "SOrL _"]) +next + case (SOrR v i \ \) + then show ?case + by (auto intro: exI[of _ "SOrR _"]) +next + case (VOr v i \ \) + then show ?case + by (auto 0 3 intro: exI[of _ "VOr _ _"]) +next + case (SAnd v i \ \) + then show ?case + by (auto 0 3 intro: exI[of _ "SAnd _ _"]) +next + case (VAndL v i \ \) + then show ?case + by (auto intro: exI[of _ "VAndL _"]) +next + case (VAndR v i \ \) + then show ?case + by (auto intro: exI[of _ "VAndR _"]) +next + case (SImpL v i \ \) + then show ?case + by (auto intro: exI[of _ "SImpL _"]) +next + case (SImpR v i \ \) + then show ?case + by (auto intro: exI[of _ "SImpR _"]) +next + case (VImp v i \ \) + then show ?case + by (auto 0 3 intro: exI[of _ "VImp _ _"]) +next + case (SIffSS v i \ \) + then show ?case + by (auto 0 3 intro: exI[of _ "SIffSS _ _"]) +next + case (SIffVV v i \ \) + then show ?case + by (auto 0 3 intro: exI[of _ "SIffVV _ _"]) +next + case (VIffSV v i \ \) + then show ?case + by (auto 0 3 intro: exI[of _ "VIffSV _ _"]) +next + case (VIffVS v i \ \) + then show ?case + by (auto 0 3 intro: exI[of _ "VIffVS _ _"]) +next + case (SExists v x i \) + then show ?case + by (auto 0 3 simp: fun_upd_def intro: exI[of _ "SExists x _ _"]) +next + case (VExists v x i \) + show ?case + proof + assume "future_bounded (\\<^sub>Fx. \)" + then have fb: "future_bounded \" + by simp + obtain mypick where mypick_def: "v_at (mypick z) = i \ v_check (v(x:=z)) \ (mypick z)" for z + using VExists fb by metis + define mypart where "mypart = tabulate (sorted_list_of_set (AD \ i)) mypick (mypick (SOME z. z \ (AD \ i)))" + have mypick_at: "\z. v_at (mypick z) = i" + by (simp add: mypick_def) + have mypick_v_check: "\z. v_check (v(x:=z)) \ (mypick z)" + by (simp add: mypick_def) + have mypick_v_check2: "\z. v_check (v(x := (SOME z. z \ AD \ i))) \ (mypick (SOME z. z \ AD \ i))" + by (simp add: mypick_def) + have v_at_myp: "v_at (VExists x mypart) = i" + using v_at_part_hd_tabulate[OF fb, of mypick i] + by (simp add: mypart_def mypick_def) + have v_check_myp: "v_check v (\\<^sub>Fx. \) (VExists x mypart)" + using v_at_tabulate[of mypick i _ \, OF mypick_at] + v_check_tabulate[OF fb mypick_at mypick_v_check] + by (auto simp add: mypart_def v_at_part_hd_tabulate[OF fb mypick_at]) + show "\vp. v_at vp = i \ v_check v (\\<^sub>Fx. \) vp" + using v_at_myp v_check_myp by blast + qed +next + case (SForall v x i \) + show ?case + proof + assume "future_bounded (\\<^sub>Fx. \)" + then have fb: "future_bounded \" + by simp + obtain mypick where mypick_def: "s_at (mypick z) = i \ s_check (v(x:=z)) \ (mypick z)" for z + using SForall fb by metis + define mypart where "mypart = tabulate (sorted_list_of_set (AD \ i)) mypick (mypick (SOME z. z \ (AD \ i)))" + have mypick_at: "\z. s_at (mypick z) = i" + by (simp add: mypick_def) + have mypick_s_check: "\z. s_check (v(x:=z)) \ (mypick z)" + by (simp add: mypick_def) + have mypick_s_check2: "\z. s_check (v(x := (SOME z. z \ AD \ i))) \ (mypick (SOME z. z \ AD \ i))" + by (simp add: mypick_def) + have s_at_myp: "s_at (SForall x mypart) = i" + using s_at_part_hd_tabulate[OF fb, of mypick i] + by (simp add: mypart_def mypick_def) + have s_check_myp: "s_check v (\\<^sub>Fx. \) (SForall x mypart)" + using s_at_tabulate[of mypick i _ \, OF mypick_at] + s_check_tabulate[OF fb mypick_at mypick_s_check] + by (auto simp add: mypart_def s_at_part_hd_tabulate[OF fb mypick_at]) + show "\sp. s_at sp = i \ s_check v (\\<^sub>Fx. \) sp" + using s_at_myp s_check_myp by blast + qed +next + case (VForall v x i \) + then show ?case + by (auto 0 3 simp: fun_upd_def intro: exI[of _ "VForall x _ _"]) +next + case (SPrev i I v \) + then show ?case + by (force intro: exI[of _ "SPrev _"]) +next + case (VPrev i v \ I) + then show ?case + by (force intro: exI[of _ "VPrev _"]) +next + case (VPrevZ i v I \) + then show ?case + by (auto intro!: exI[of _ "VPrevZ"]) +next + case (VPrevOutL i I v \) + then show ?case + by (auto intro!: exI[of _ "VPrevOutL i"]) +next + case (VPrevOutR i I v \) + then show ?case + by (auto intro!: exI[of _ "VPrevOutR i"]) +next + case (SNext i I v \) + then show ?case + by (force simp: Let_def intro: exI[of _ "SNext _"]) +next + case (VNext v i \ I) + then show ?case + by (force simp: Let_def intro: exI[of _ "VNext _"]) +next + case (VNextOutL i I v \) + then show ?case + by (auto intro!: exI[of _ "VNextOutL i"]) +next + case (VNextOutR i I v \) + then show ?case + by (auto intro!: exI[of _ "VNextOutR i"]) +next + case (SOnce j i I v \) + then show ?case + by (auto simp: Let_def intro: exI[of _ "SOnce i _"]) +next + case (VOnceOut i I v \) + then show ?case + by (auto intro!: exI[of _ "VOnceOut i"]) +next + case (VOnce j I i v \) + show ?case + proof + assume "future_bounded (\<^bold>P I \)" + then have fb: "future_bounded \" + by simp + obtain mypick where mypick_def: "\k \ {j .. LTP_p \ i I}. v_at (mypick k) = k \ v_check v \ (mypick k)" + using VOnce fb by metis + then obtain vps where vps_def: "map (v_at) vps = [j ..< Suc (LTP_p \ i I)] \ (\vp \ set vps. v_check v \ vp)" + by atomize_elim (auto intro!: trans[OF list.map_cong list.map_id] exI[of _ "map mypick ([j ..< Suc (LTP_p \ i I)])"]) + then have "v_at (VOnce i j vps) = i \ v_check v (\<^bold>P I \) (VOnce i j vps)" + using VOnce by auto + then show "\vp. v_at vp = i \ v_check v (\<^bold>P I \) vp" + by blast + qed +next + case (SEventually j i I v \) + then show ?case + by (auto simp: Let_def intro: exI[of _ "SEventually i _"]) +next + case (VEventually I i v \) + show ?case + proof + assume fb_eventually: "future_bounded (\<^bold>F I \)" + then have fb: "future_bounded \" + by simp + obtain b where b_def: "right I = enat b" + using fb_eventually by (atomize_elim, cases "right I") auto + define j where j_def: "j = LTP \ (\ \ i + b)" + obtain mypick where mypick_def: "\k \ {ETP_f \ i I .. j}. v_at (mypick k) = k \ v_check v \ (mypick k)" + using VEventually fb_eventually unfolding b_def j_def enat.simps + by atomize_elim (rule bchoice, simp) + then obtain vps where vps_def: "map (v_at) vps = [ETP_f \ i I ..< Suc j] \ (\vp \ set vps. v_check v \ vp)" + by atomize_elim (auto intro!: trans[OF list.map_cong list.map_id] exI[of _ "map mypick ([ETP_f \ i I ..< Suc j])"]) + then have "v_at (VEventually i j vps) = i \ v_check v (\<^bold>F I \) (VEventually i j vps)" + using VEventually b_def j_def by simp + then show "\vp. v_at vp = i \ v_check v (\<^bold>F I \) vp" + by blast + qed +next + case (SHistorically j I i v \) + show ?case + proof + assume fb_historically: "future_bounded (\<^bold>H I \)" + then have fb: "future_bounded \" + by simp + obtain mypick where mypick_def: "\k \ {j .. LTP_p \ i I}. s_at (mypick k) = k \ s_check v \ (mypick k)" + using SHistorically fb by metis + then obtain sps where sps_def: "map (s_at) sps = [j ..< Suc (LTP_p \ i I)] \ (\sp \ set sps. s_check v \ sp)" + by atomize_elim (auto intro!: trans[OF list.map_cong list.map_id] exI[of _ "map mypick ([j ..< Suc (LTP_p \ i I)])"]) + then have "s_at (SHistorically i j sps) = i \ s_check v (\<^bold>H I \) (SHistorically i j sps)" + using SHistorically by auto + then show "\sp. s_at sp = i \ s_check v (\<^bold>H I \) sp" + by blast + qed +next + case (SHistoricallyOut i I v \) + then show ?case + by (auto intro!: exI[of _ "SHistoricallyOut i"]) +next + case (VHistorically j i I v \) + then show ?case + by (auto simp: Let_def intro: exI[of _ "VHistorically i _"]) +next + case (SAlways I i v \) + show ?case + proof + assume fb_always: "future_bounded (\<^bold>G I \)" + then have fb: "future_bounded \" + by simp + obtain b where b_def: "right I = enat b" + using fb_always by (atomize_elim, cases "right I") auto + define j where j_def: "j = LTP \ (\ \ i + b)" + obtain mypick where mypick_def: "\k \ {ETP_f \ i I .. j}. s_at (mypick k) = k \ s_check v \ (mypick k)" + using SAlways fb_always unfolding b_def j_def enat.simps + by atomize_elim (rule bchoice, simp) + then obtain sps where sps_def: "map (s_at) sps = [ETP_f \ i I ..< Suc j] \ (\sp \ set sps. s_check v \ sp)" + by atomize_elim (auto intro!: trans[OF list.map_cong list.map_id] exI[of _ "map mypick ([ETP_f \ i I ..< Suc j])"]) + then have "s_at (SAlways i j sps) = i \ s_check v (\<^bold>G I \) (SAlways i j sps)" + using SAlways b_def j_def by simp + then show "\sp. s_at sp = i \ s_check v (\<^bold>G I \) sp" + by blast + qed +next + case (VAlways j i I v \) + then show ?case + by (auto simp: Let_def intro: exI[of _ "VAlways i _"]) +next + case (SSince j i I v \ \) + show ?case + proof + assume fb_since: "future_bounded (\ \<^bold>S I \)" + then have fb: "future_bounded \" "future_bounded \" + by simp_all + obtain sp2 where sp2_def: "s_at sp2 = j \ s_check v \ sp2" + using SSince fb_since by auto + { + assume "Suc j > i" + then have "s_at (SSince sp2 []) = i \ s_check v (\ \<^bold>S I \) (SSince sp2 [])" + using sp2_def SSince by auto + then have "\sp. s_at sp = i \ s_check v (\ \<^bold>S I \) sp" + by blast + } + moreover + { + assume sucj_leq_i: "Suc j \ i" + obtain mypick where mypick_def: "\k \ {Suc j ..< Suc i}. s_at (mypick k) = k \ s_check v \ (mypick k)" + using SSince fb_since by atomize_elim (rule bchoice, simp) + then obtain sp1s where sp1s_def: "map (s_at) sp1s = [Suc j ..< Suc i] \ (\sp \ set sp1s. s_check v \ sp)" + by atomize_elim (auto intro!: trans[OF list.map_cong list.map_id] exI[of _ "map mypick ([Suc j ..< Suc i])"]) + then have "sp1s \ []" + using sucj_leq_i by auto + then have "s_at (SSince sp2 sp1s) = i \ s_check v (\ \<^bold>S I \) (SSince sp2 sp1s)" + using SSince sucj_leq_i fb sp2_def sp1s_def + by (clarsimp simp add: + Cons_eq_upt_conv append_eq_Cons_conv map_eq_append_conv + split: list.splits) auto + then have "\sp. s_at sp = i \ s_check v (\ \<^bold>S I \) sp" + by blast + } + ultimately show "\sp. s_at sp = i \ s_check v (\ \<^bold>S I \) sp" + using not_less by blast + qed +next + case (VSinceOut i I v \ \) + then show ?case + by (auto intro!: exI[of _ "VSinceOut i"]) +next + case (VSince I i j v \ \) + show ?case + proof + assume fb_since: "future_bounded (\ \<^bold>S I \)" + then have fb: "future_bounded \" "future_bounded \" + by simp_all + obtain vp1 where vp1_def: "v_at vp1 = j \ v_check v \ vp1" + using fb_since VSince by auto + obtain mypick where mypick_def: "\k \ {j .. LTP_p \ i I}. v_at (mypick k) = k \ v_check v \ (mypick k)" + using VSince fb_since by atomize_elim (rule bchoice, simp) + then obtain vp2s where vp2s_def: "map (v_at) vp2s = [j ..< Suc (LTP_p \ i I)] \ (\vp \ set vp2s. v_check v \ vp)" + by atomize_elim (auto intro!: trans[OF list.map_cong list.map_id] exI[of _ "map mypick ([j ..< Suc (LTP_p \ i I)])"]) + then have "v_at (VSince i vp1 vp2s) = i \ v_check v (\ \<^bold>S I \) (VSince i vp1 vp2s)" + using vp1_def VSince by auto + then show "\vp. v_at vp = i \ v_check v (\ \<^bold>S I \) vp" + by blast + qed +next + case (VSinceInf j I i v \ \) + show ?case + proof + assume fb_since: "future_bounded (\ \<^bold>S I \)" + then have fb: "future_bounded \" "future_bounded \" + by simp_all + obtain mypick where mypick_def: "\k \ {j .. LTP_p \ i I}. v_at (mypick k) = k \ v_check v \ (mypick k)" + using VSinceInf fb_since by atomize_elim (rule bchoice, simp) + then obtain vp2s where vp2s_def: "map (v_at) vp2s = [j ..< Suc (LTP_p \ i I)] \ (\vp \ set vp2s. v_check v \ vp)" + by atomize_elim (auto intro!: trans[OF list.map_cong list.map_id] exI[of _ "map mypick ([j ..< Suc (LTP_p \ i I)])"]) + then have "v_at (VSinceInf i j vp2s) = i \ v_check v (\ \<^bold>S I \) (VSinceInf i j vp2s)" + using VSinceInf by auto + then show "\vp. v_at vp = i \ v_check v (\ \<^bold>S I \) vp" + by blast + qed +next + case (SUntil j i I v \ \) + show ?case + proof + assume fb_until: "future_bounded (\ \<^bold>U I \)" + then have fb: "future_bounded \" "future_bounded \" + by simp_all + obtain sp2 where sp2_def: "s_at sp2 = j \ s_check v \ sp2" + using fb SUntil by blast + { + assume "i \ j" + then have "s_at (SUntil [] sp2) = i \ s_check v (\ \<^bold>U I \) (SUntil [] sp2)" + using sp2_def SUntil by auto + then have "\sp. s_at sp = i \ s_check v (\ \<^bold>U I \) sp" + by blast + } + moreover + { + assume i_l_j: "i < j" + obtain mypick where mypick_def: "\k \ {i ..< j}. s_at (mypick k) = k \ s_check v \ (mypick k)" + using SUntil fb_until by atomize_elim (rule bchoice, simp) + then obtain sp1s where sp1s_def: "map (s_at) sp1s = [i ..< j] \ (\sp \ set sp1s. s_check v \ sp)" + by atomize_elim (auto intro!: trans[OF list.map_cong list.map_id] exI[of _ "map mypick ([i ..< j])"]) + then have "s_at (SUntil sp1s sp2) = i \ s_check v (\ \<^bold>U I \) (SUntil sp1s sp2)" + using SUntil fb_until sp2_def sp1s_def i_l_j + by (clarsimp simp add: append_eq_Cons_conv map_eq_append_conv split: list.splits) + (auto simp: Cons_eq_upt_conv dest!: upt_eq_Nil_conv[THEN iffD1, OF sym]) + then have "\sp. s_at sp = i \ s_check v (\ \<^bold>U I \) sp" + by blast + } + ultimately show "\sp. s_at sp = i \ s_check v (\ \<^bold>U I \) sp" + using not_less by blast + qed +next + case (VUntil I j i v \ \) + show ?case + proof + assume fb_until: "future_bounded (\ \<^bold>U I \)" + then have fb: "future_bounded \" "future_bounded \" + by simp_all + obtain vp1 where vp1_def: "v_at vp1 = j \ v_check v \ vp1" + using VUntil fb_until by auto + obtain mypick where mypick_def: "\k \ {ETP_f \ i I .. j}. v_at (mypick k) = k \ v_check v \ (mypick k)" + using VUntil fb_until by atomize_elim (rule bchoice, simp) + then obtain vp2s where vp2s_def: "map (v_at) vp2s = [ETP_f \ i I ..< Suc j] \ (\vp \ set vp2s. v_check v \ vp)" + by atomize_elim (auto intro!: trans[OF list.map_cong list.map_id] exI[of _ "map mypick ([ETP_f \ i I ..< Suc j])"]) + then have "v_at (VUntil i vp2s vp1) = i \ v_check v (\ \<^bold>U I \) (VUntil i vp2s vp1)" + using VUntil fb_until vp1_def by simp + then show "\vp. v_at vp = i \ v_check v (\ \<^bold>U I \) vp" + by blast + qed +next + case (VUntilInf I i v \ \) + show ?case + proof + assume fb_until: "future_bounded (\ \<^bold>U I \)" + then have fb: "future_bounded \" "future_bounded \" + by simp_all + obtain b where b_def: "right I = enat b" + using fb_until by (atomize_elim, cases "right I") auto + define j where j_def: "j = LTP \ (\ \ i + b)" + obtain mypick where mypick_def: "\k \ {ETP_f \ i I .. j}. v_at (mypick k) = k \ v_check v \ (mypick k)" + using VUntilInf fb_until unfolding b_def j_def by atomize_elim (rule bchoice, simp) + then obtain vp2s where vp2s_def: "map (v_at) vp2s = [ETP_f \ i I ..< Suc j] \ (\vp \ set vp2s. v_check v \ vp)" + by atomize_elim (auto intro!: trans[OF list.map_cong list.map_id] exI[of _ "map mypick ([ETP_f \ i I ..< Suc j])"]) + then have "v_at (VUntilInf i j vp2s) = i \ v_check v (\ \<^bold>U I \) (VUntilInf i j vp2s)" + using VUntilInf b_def j_def by simp + then show "\vp. v_at vp = i \ v_check v (\ \<^bold>U I \) vp" + by blast + qed +qed + +lemmas check_completeness = + conjunct1[OF check_completeness_aux, rule_format] + conjunct2[OF check_completeness_aux, rule_format] + +definition "p_check v \ p = (case p of Inl sp \ s_check v \ sp | Inr vp \ v_check v \ vp)" +definition "p_check_exec vs \ p = (case p of Inl sp \ s_check_exec vs \ sp | Inr vp \ v_check_exec vs \ vp)" + +definition valid :: "('n, 'd) envset \ nat \ ('n, 'd) formula \ ('n, 'd) proof \ bool" where + "valid vs i \ p = + (case p of + Inl p \ s_check_exec vs \ p \ s_at p = i + | Inr p \ v_check_exec vs \ p \ v_at p = i)" + +end + +subsection \Lifting the Checker to PDTs\ + +fun check_one where + "check_one \ v \ (Leaf p) = p_check \ v \ p" +| "check_one \ v \ (Node x part) = check_one \ v \ (lookup_part part (v x))" + +fun check_all_aux where + "check_all_aux \ vs \ (Leaf p) = p_check_exec \ vs \ p" +| "check_all_aux \ vs \ (Node x part) = (\(D, e) \ set (subsvals part). check_all_aux \ (vs(x := D)) \ e)" + +fun collect_paths_aux where + "collect_paths_aux DS \ vs \ (Leaf p) = (if p_check_exec \ vs \ p then {} else rev ` DS)" +| "collect_paths_aux DS \ vs \ (Node x part) = (\(D, e) \ set (subsvals part). collect_paths_aux (Cons D ` DS) \ (vs(x := D)) \ e)" + +lemma check_one_cong: "\x\fv \ \ vars e. v x = v' x \ check_one \ v \ e = check_one \ v' \ e" +proof (induct e arbitrary: v v') + case (Leaf x) + then show ?case + by (auto simp: p_check_def check_fv_cong split: sum.splits) +next + case (Node x part) + from Node(2) have *: "v x = v' x" + by simp + from Node(2) show ?case + unfolding check_one.simps * + by (intro Node(1)) auto +qed + +lemma check_all_aux_check_one: "\x. vs x \ {} \ distinct_paths e \ (\x \ vars e. vs x = UNIV) \ + check_all_aux \ vs \ e \ (\v \ compatible_vals (fv \) vs. check_one \ v \ e)" +proof (induct e arbitrary: vs) + case (Node x part) + show ?case + unfolding check_all_aux.simps check_one.simps split_beta + proof (safe, unfold fst_conv snd_conv, goal_cases LR RL) + case (LR v) + from Node(2-) fst_lookup[of "v x" part] LR(1)[rule_format, OF lookup_subsvals[of _ "v x"]] LR(2) show ?case + by (subst (asm) Node(1)) + (auto 0 3 simp: compatible_vals_fun_upd dest!: bspec[of _ _ v] + elim!: compatible_vals_antimono[THEN set_mp, rotated]) + next + case (RL D e) + from RL(2) obtain d where "d \ D" + by transfer (force simp: partition_on_def image_iff) + with RL show ?case + using Node(2-) lookup_subsvals[of part d] lookup_part_Vals[of part d] + lookup_part_from_subvals[of D e part d] + proof (intro Node(1)[THEN iffD2, OF _ _ _ _ ballI], goal_cases _ _ _ _ compatible) + case (compatible v) + from compatible(2-) compatible(1)[THEN bspec, of "v(x := d)"] compatible(1)[THEN bspec, of v] + show ?case + using lookup_part_from_subvals[of D e part "v x"] + fun_upd_in_compatible_vals_in[of v "fv \" x vs "v x"] + check_one_cong[THEN iffD1, rotated -1, of \ "v(x := d)" \ e v, simplified] + by (auto simp: compatible_vals_fun_upd fun_upd_apply[of _ _ _ x] + fun_upd_in_compatible_vals_notin split: if_splits + simp del: fun_upd_apply) + qed auto + qed +qed (auto simp: p_check_exec_def p_check_def check_exec_check split: sum.splits) + +definition check_all :: "('n, 'd :: {default, linorder}) trace \ ('n, 'd) formula \ ('n, 'd) expl \ bool" where + "check_all \ \ e = (distinct_paths e \ check_all_aux \ (\_. UNIV) \ e)" + +lemma check_one_alt: "check_one \ v \ e = p_check \ v \ (eval_pdt v e)" + by (induct e) auto + +lemma check_all_alt: "check_all \ \ e = (distinct_paths e \ (\v. p_check \ v \ (eval_pdt v e)))" + unfolding check_all_def + by (rule conj_cong[OF refl], subst check_all_aux_check_one) + (auto simp: compatible_vals_def check_one_alt) + +fun pdt_at where + "pdt_at i (Leaf l) = (p_at l = i)" +| "pdt_at i (Node x part) = (\pdt \ Vals part. pdt_at i pdt)" + +lemma pdt_at_p_at_eval_pdt: "pdt_at i e \ p_at (eval_pdt v e) = i" + by (induct e) auto + +lemma check_all_completeness_aux: + fixes \ :: "('n, 'd :: {default, linorder}) formula" + shows "set vs \ fv \ \ future_bounded \ \ distinct vs \ + \e. pdt_at i e \ vars_order vs e \ (\v. (\x. x \ set vs \ v x = w x) \ p_check \ v \ (eval_pdt v e))" +proof (induct vs arbitrary: w) + case Nil + then show ?case + proof (cases "sat \ w i \") + case True + then have "SAT \ w i \" by (rule completeness) + with Nil obtain sp where "s_at sp = i" "s_check \ w \ sp" by (blast dest: check_completeness) + then show ?thesis + by (intro exI[of _ "Leaf (Inl sp)"]) (auto simp: vars_order.intros p_check_def p_at_def) + next + case False + then have "VIO \ w i \" by (rule completeness) + with Nil obtain vp where "v_at vp = i" "v_check \ w \ vp" by (blast dest: check_completeness) + then show ?thesis + by (intro exI[of _ "Leaf (Inr vp)"]) (auto simp: vars_order.intros p_check_def p_at_def) + qed +next + case (Cons x vs) + define eq :: "('n \ 'd) \ ('n \ 'd) \ bool" where "eq = rel_fun (eq_onp (\x. x \ set vs)) (=)" + from Cons have "\w. \e. pdt_at i e \ vars_order vs e \ + (\v. (\x. x \ set vs \ v x = w x) \ p_check \ v \ (eval_pdt v e))" by simp + then obtain pick :: "'d \ ('n, 'd) expl" where pick: "pdt_at i (pick a)" "vars_order vs (pick a)" and + eq_pick: "\v. eq v (w(x := a)) \ p_check \ v \ (eval_pdt v (pick a))" for a + unfolding eq_def rel_fun_def eq_onp_def choice_iff + proof (atomize_elim, elim exE, goal_cases pick_val) + case (pick_val f) + then show ?case + by (auto intro!: exI[of _ "\a. f (w(x := a))"]) + qed + let ?a = "SOME z. z \ AD \ \ i" + let ?AD = "sorted_list_of_set (AD \ \ i)" + show ?case + proof (intro exI[of _ "Node x (tabulate ?AD pick (pick ?a))"] conjI allI impI, + goal_cases pdt_at vars_order p_check) + case (p_check w') + have "w' x \ AD \ \ i \ ?a \ AD \ \ i" + by (metis some_eq_imp) + moreover have "eq (w'(x := ?a)) (w(x := ?a))" + using p_check by (auto simp: eq_def rel_fun_def eq_onp_def) + moreover have "eq w' (w(x := w' x))" + using p_check by (auto simp: eq_def rel_fun_def eq_onp_def) + ultimately show ?case + using pick Cons(2-) eq_pick[of w' "w' x"] eq_pick[of "w'(x := ?a)" ?a] + pdt_at_p_at_eval_pdt[of "i" "pick ?a" w'] eval_pdt_fun_upd[of vs "pick ?a" x w' ?a] + by (auto simp: p_check_def p_at_def + elim!: check_AD_cong[THEN iffD1, rotated -1, of _ _ _ _ _ i] + split: if_splits sum.splits sum.splits) + qed (use Cons(2-) pick in \simp_all add: vars_order.intros\) +qed + +lemma check_all_completeness: + fixes \ :: "('n, 'd :: {default, linorder}) formula" + assumes "future_bounded \" + shows "\e. pdt_at i e \ check_all \ \ e" +proof - + obtain vs where vs[simp]: "distinct vs" "set vs = fv \" + by (meson finite_distinct_list finite_fv) + have s: "s_check \ v \ sp" + if "vars_order vs e" + and "\v. (\sp. eval_pdt v e = Inl sp \ (\x. x \ fv \ \ v x \ undefined) \ s_check \ v \ sp) \ + (\vp. eval_pdt v e = Inr vp \ (\x. x \ fv \ \ v x \ undefined) \ v_check \ v \ vp)" + and "eval_pdt v e = Inl sp" for e v sp + using that eval_pdt_cong[of e v "\x. if x \ fv \ then v x else undefined"] + check_fv_cong[of \ v "\x. if x \ fv \ then v x else undefined"] + by (auto dest!: spec[of _ sp] vars_order_vars simp: subset_eq) + have v: "v_check \ v \ vp" + if "vars_order vs e" + and "\v. (\sp. eval_pdt v e = Inl sp \ (\x. x \ fv \ \ v x \ undefined) \ s_check \ v \ sp) \ + (\vp. eval_pdt v e = Inr vp \ (\x. x \ fv \ \ v x \ undefined) \ v_check \ v \ vp)" + and "eval_pdt v e = Inr vp" for e v vp + using that eval_pdt_cong[of e v "\x. if x \ fv \ then v x else undefined"] + check_fv_cong[of \ v "\x. if x \ fv \ then v x else undefined"] + by (auto dest!: spec[of _ vp] vars_order_vars simp: subset_eq) + show ?thesis + using check_all_completeness_aux[of vs \ i "\_. undefined" \] assms + unfolding check_all_alt p_check_def + by (auto elim!: exI [where P = "\x. _ x \ _ x" , OF conjI] simp: vars_order_distinct_paths split: sum.splits intro: s v) +qed + +lemma check_all_soundness_aux: "check_all \ \ e \ p = eval_pdt v e \ isl p \ sat \ v (p_at p) \" + unfolding check_all_alt + by (auto simp: isl_def p_check_def p_at_def dest!: spec[of _ v] + dest: check_soundness soundness split: sum.splits) + +lemma check_all_soundness: "check_all \ \ e \ pdt_at i e \ isl (eval_pdt v e) \ sat \ v i \" + by (drule check_all_soundness_aux[OF _ refl, of _ _ _ v]) (auto simp: pdt_at_p_at_eval_pdt) + +unbundle MFOTL_no_notation \ \ disable notation \ + +(*<*) +end +(*>*) \ No newline at end of file diff --git a/thys/MFOTL_Checker/Checker_Code.thy b/thys/MFOTL_Checker/Checker_Code.thy new file mode 100644 --- /dev/null +++ b/thys/MFOTL_Checker/Checker_Code.thy @@ -0,0 +1,868 @@ +(*<*) +theory Checker_Code + imports + Checker Event_Data + "HOL-Library.Code_Target_Nat" + "HOL.String" + "HOL-Library.List_Lexorder" + "HOL-Library.AList_Mapping" + "HOL-Library.Cardinality" +begin +(*>*) + +section \Code Generation\ + +subsection \Type Class Instances\ + +class universe = + fixes universe :: "'a list option" + assumes infinite: "universe = None \ infinite (UNIV :: 'a set)" + and finite: "universe = Some xs \ distinct xs \ set xs = UNIV" +begin + +lemma finite_coset: "finite (List.coset (xs :: 'a list)) = (case universe of None \ False | _ \ True)" + using infinite finite + by (auto split: option.splits dest!: equalityD2 elim!: finite_subset) + +end + +declare [[code drop: finite]] +declare finite_set[THEN eqTrueI, code] finite_coset[code] + +instantiation bool :: universe begin +definition universe_bool :: "bool list option" where "universe_bool = Some [True, False]" +instance by standard (auto simp: universe_bool_def) +end +instantiation char :: universe begin +definition universe_char :: "char list option" where "universe_char = Some (map char_of [0::nat..<256])" +instance by standard (use UNIV_char_of_nat in \auto simp: universe_char_def distinct_map\) +end +instantiation nat :: universe begin +definition universe_nat :: "nat list option" where "universe_nat = None" +instance by standard (auto simp: universe_nat_def) +end +instantiation list :: (type) universe begin +definition universe_list :: "'a list list option" where "universe_list = None" +instance by standard (auto simp: universe_list_def infinite_UNIV_listI) +end +instantiation String.literal :: universe begin +definition universe_literal :: "String.literal list option" where "universe_literal = None" +instance by standard (auto simp: universe_literal_def infinite_literal) +end +instantiation string8 :: universe begin +definition universe_string8 :: "string8 list option" where "universe_string8 = None" +lemma UNIV_string8: "UNIV = Abs_string8 ` UNIV" + by (auto simp: image_iff intro: Abs_string8_cases) +instance by standard + (auto simp: universe_string8_def UNIV_string8 finite_image_iff Abs_string8_inject inj_on_def infinite_UNIV_listI) +end +instantiation prod :: (universe, universe) universe begin +definition universe_prod :: "('a \ 'b) list option" where "universe_prod = + (case (universe, universe) of (Some xs, Some ys) \ Some (List.product xs ys) | _ \ None)" +instance by standard + (auto simp: universe_prod_def finite_prod distinct_product infinite finite split: option.splits) +end +instantiation sum :: (universe, universe) universe begin +definition universe_sum :: "('a + 'b) list option" where "universe_sum = + (case (universe, universe) of (Some xs, Some ys) \ Some (map Inl xs @ map Inr ys) | _ \ None)" +instance by standard + (use UNIV_sum in \auto simp: universe_sum_def distinct_map infinite finite split: option.splits\) +end +instantiation option :: (universe) universe begin +definition "universe_option = (case universe of Some xs \ Some (None # map Some xs) | _ \ None)" +instance by standard (auto simp: universe_option_def distinct_map finite infinite image_iff split: option.splits) +end +instantiation "fun" :: (universe, universe) universe begin +definition universe_fun :: "('a \ 'b) list option" where "universe_fun = + (case (universe, universe) of + (Some xs, Some ys) \ Some (map (\zs. the \ map_of (zip xs zs)) (List.n_lists (length xs) ys)) + | (_, Some [x]) \ Some [\_. x] + | _ \ None)" +instance +proof - + have 1: False if "infinite (UNIV::'a set)" "CARD('b) = Suc 0" "a \ b" for a b :: 'b + using that by (metis (full_types) UNIV_I card_1_singleton_iff singletonD) + have 2: "ys = zs" + if "distinct (xs::'a list)" and "length ys = length xs" and "length zs = length xs" + and "\a. the (map_of (zip xs ys) a) = the (map_of (zip xs zs) a)" + for xs :: "'a list" and ys zs :: "'b list" + using that by (metis map_fst_zip map_of_eqI map_of_zip_inject map_of_zip_is_None option.expand) + have 3: "\zs. length zs = length xs \ set zs \ set ys \ (\x. f x = the (map_of (zip xs zs) x))" + if "\x. x \ set xs" "\x. x \ set ys" + for xs ys and f :: "'a \ 'b" + using that by (metis length_map map_of_zip_map option.sel subsetI) + show "OFCLASS('a \ 'b, universe_class)" + by standard + (auto 0 3 simp: universe_fun_def set_eq_iff fun_eq_iff image_iff set_n_lists distinct_map + inj_on_def distinct_n_lists finite_UNIV_fun dest!: infinite finite + split: option.splits list.splits intro: 1 2 3) +qed +end +instantiation event_data :: universe begin +definition universe_event_data :: "event_data list option" where "universe_event_data = None" +instance by standard (simp_all add: infinite_UNIV_event_data universe_event_data_def) +end + +instantiation nat :: default begin +definition default_nat :: nat where "default_nat = 0" +instance proof qed +end + +instantiation list :: (type) default begin +definition default_list :: "'a list" where "default_list = []" +instance proof qed +end + +instance event_data :: equal by standard + +instantiation String.literal :: default begin +definition default_literal :: String.literal where "default_literal = 0" +instance proof qed +end + +instantiation event_data :: card_UNIV begin +definition "finite_UNIV = Phantom(event_data) False" +definition "card_UNIV = Phantom(event_data) 0" +instance by intro_classes (simp_all add: finite_UNIV_event_data_def card_UNIV_event_data_def infinite_UNIV_event_data) +end + +subsection \Progress\ + +primrec progress :: "('n, 'd) trace \ ('n, 'd) Formula.formula \ nat \ nat" where + "progress \ Formula.TT j = j" +| "progress \ Formula.FF j = j" +| "progress \ (Formula.Eq_Const _ _) j = j" +| "progress \ (Formula.Pred _ _) j = j" +| "progress \ (Formula.Neg \) j = progress \ \ j" +| "progress \ (Formula.Or \ \) j = min (progress \ \ j) (progress \ \ j)" +| "progress \ (Formula.And \ \) j = min (progress \ \ j) (progress \ \ j)" +| "progress \ (Formula.Imp \ \) j = min (progress \ \ j) (progress \ \ j)" +| "progress \ (Formula.Iff \ \) j = min (progress \ \ j) (progress \ \ j)" +| "progress \ (Formula.Exists _ \) j = progress \ \ j" +| "progress \ (Formula.Forall _ \) j = progress \ \ j" +| "progress \ (Formula.Prev I \) j = (if j = 0 then 0 else min (Suc (progress \ \ j)) j)" +| "progress \ (Formula.Next I \) j = progress \ \ j - 1" +| "progress \ (Formula.Once I \) j = progress \ \ j" +| "progress \ (Formula.Historically I \) j = progress \ \ j" +| "progress \ (Formula.Eventually I \) j = + Inf {i. \k. k < j \ k \ (progress \ \ j) \ (\ \ k - \ \ i) \ right I}" +| "progress \ (Formula.Always I \) j = + Inf {i. \k. k < j \ k \ (progress \ \ j) \ (\ \ k - \ \ i) \ right I}" +| "progress \ (Formula.Since \ I \) j = min (progress \ \ j) (progress \ \ j)" +| "progress \ (Formula.Until \ I \) j = + Inf {i. \k. k < j \ k \ min (progress \ \ j) (progress \ \ j) \ (\ \ k - \ \ i) \ right I}" + +lemma Inf_Min: + fixes P :: "nat \ bool" + assumes "P j" + shows "Inf (Collect P) = Min (Set.filter P {..j})" + using Min_in[where ?A="Set.filter P {..j}"] assms + by (auto simp: Set.filter_def intro: cInf_lower intro!: antisym[OF _ Min_le]) + (metis Inf_nat_def1 empty_iff mem_Collect_eq) + +lemma progress_Eventually_code: "progress \ (Formula.Eventually I \) j = + (let m = min j (Suc (progress \ \ j)) - 1 in Min (Set.filter (\i. enat (\ \ m i) \ right I) {..j}))" +proof - + define P where "P \ (\i. \k. k < j \ k \ (progress \ \ j) \ enat (\ \ k i) \ right I)" + have P_j: "P j" + by (auto simp: P_def enat_0) + have all_wit: "(\k \ {.. \ k i) \ right I) \ (enat (\ \ (m - 1) i) \ right I)" for i m + proof (cases m) + case (Suc ma) + have "k < Suc ma \ \ \ k i \ \ \ ma i" for k + using \_mono + unfolding less_Suc_eq_le + by (rule diff_le_mono) + then show ?thesis + by (auto simp: Suc) (meson order.trans enat_ord_simps(1)) + qed (auto simp: enat_0) + show ?thesis + unfolding progress.simps Let_def P_def[symmetric] Inf_Min[where ?P=P, OF P_j] all_wit[symmetric] + by (fastforce simp: P_def intro: arg_cong[where ?f=Min]) +qed + +lemma progress_Always_code: "progress \ (Formula.Always I \) j = + (let m = min j (Suc (progress \ \ j)) - 1 in Min (Set.filter (\i. enat (\ \ m i) \ right I) {..j}))" +proof - + define P where "P \ (\i. \k. k < j \ k \ (progress \ \ j) \ enat (\ \ k i) \ right I)" + have P_j: "P j" + by (auto simp: P_def enat_0) + have all_wit: "(\k \ {.. \ k i) \ right I) \ (enat (\ \ (m - 1) i) \ right I)" for i m + proof (cases m) + case (Suc ma) + have "k < Suc ma \ \ \ k i \ \ \ ma i" for k + using \_mono + unfolding less_Suc_eq_le + by (rule diff_le_mono) + then show ?thesis + by (auto simp: Suc) (meson order.trans enat_ord_simps(1)) + qed (auto simp: enat_0) + show ?thesis + unfolding progress.simps Let_def P_def[symmetric] Inf_Min[where ?P=P, OF P_j] all_wit[symmetric] + by (fastforce simp: P_def intro: arg_cong[where ?f=Min]) +qed + +lemma progress_Until_code: "progress \ (Formula.Until \ I \) j = + (let m = min j (Suc (min (progress \ \ j) (progress \ \ j))) - 1 in Min (Set.filter (\i. enat (\ \ m i) \ right I) {..j}))" +proof - + define P where "P \ (\i. \k. k < j \ k \ min (progress \ \ j) (progress \ \ j) \ enat (\ \ k i) \ right I)" + have P_j: "P j" + by (auto simp: P_def enat_0) + have all_wit: "(\k \ {.. \ k i) \ right I) \ (enat (\ \ (m - 1) i) \ right I)" for i m + proof (cases m) + case (Suc ma) + have "k < Suc ma \ \ \ k i \ \ \ ma i" for k + using \_mono + unfolding less_Suc_eq_le + by (rule diff_le_mono) + then show ?thesis + by (auto simp: Suc) (meson order.trans enat_ord_simps(1)) + qed (auto simp: enat_0) + show ?thesis + unfolding progress.simps Let_def P_def[symmetric] Inf_Min[where ?P=P, OF P_j] all_wit[symmetric] + by (fastforce simp: P_def intro: arg_cong[where ?f=Min]) +qed + +lemmas progress_code[code] = progress.simps(1-15) progress_Eventually_code progress_Always_code progress.simps(18) progress_Until_code + +subsection \Trace\ + +lemma snth_Stream_eq: "(x ## s) !! n = (case n of 0 \ x | Suc m \ s !! m)" + by (cases n) auto + +lemma extend_is_stream: + assumes "sorted (map snd list)" + and "\x. x \ set list \ snd x \ m" + and "\x. x \ set list \ finite (fst x)" + shows "ssorted (smap snd (list @- smap (\n. ({}, n + m)) nats)) \ + sincreasing (smap snd (list @- smap (\n. ({}, n + m)) nats)) \ + sfinite (smap fst (list @- smap (\n. ({}, n + m)) nats))" +proof - + have A: "\x\set list. n \ snd x \ n \ m \ + n \ (map snd list @- smap (\x. x + m) nats) !! i" for n i + and list :: "('a set \ nat) list" + proof (induction i arbitrary: n) + case (Suc i) + then show ?case + by (auto simp: shift_snth nth_tl) + qed (auto simp add: list.map_sel(1)) + then have "ssorted (smap snd (list @- smap (\n. ({}, n + m)) nats))" + using assms + by (induction list) (auto simp: stream.map_comp o_def ssorted_iff_mono snth_Stream_eq + split: nat.splits) + moreover have "sincreasing (smap snd (list @- smap (\n. ({}, n + m)) nats))" + using assms + proof (induction list) + case Nil + then show ?case + by (simp add: sincreasing_def) presburger + next + case (Cons a as) + have IH: "\x. \i. x < smap snd (as @- smap (\n. ({}, n + m)) nats) !! i" + using Cons + by (simp add: sincreasing_def) + show ?case + using IH + by (simp add: sincreasing_def) + (metis snth_Stream) + qed + moreover have "sfinite (smap fst (list @- smap (\n. ({}, n + m)) nats))" + using assms(3) + proof (induction list) + case Nil + then show ?case by (simp add: sfinite_def) + next + case (Cons a as) + then have fin: "finite (fst a)" + by simp + show ?case + using Cons + by (auto simp add: sfinite_def snth_Stream_eq split: nat.splits) + qed + ultimately show ?thesis + by simp +qed + +typedef 'a trace_mapping = "{(n, m, t) :: (nat \ nat \ (nat, 'a set \ nat) mapping) | + n m t. Mapping.keys t = {.. + sorted (map (snd \ (the \ Mapping.lookup t)) [0.. + (case n of 0 \ True | Suc n' \ (case Mapping.lookup t n' of Some (X', t') \ t' \ m | None \ False)) \ + (\n' < n. case Mapping.lookup t n' of Some (X', t') \ finite X' | None \ False)}" + by (rule exI[of _ "(0, 0, Mapping.empty)"]) auto + +setup_lifting type_definition_trace_mapping + +lemma lookup_bulkload_Some: "i < length list \ + Mapping.lookup (Mapping.bulkload list) i = Some (list ! i)" + by transfer auto + +lift_definition trace_mapping_of_list :: "('a set \ nat) list \ 'a trace_mapping" is + "\xs. if sorted (map snd xs) \ (\x \ set xs. finite (fst x)) then (if xs = [] then (0, 0, Mapping.empty) + else (length xs, snd (last xs), Mapping.bulkload xs)) + else (0, 0, Mapping.empty)" + by (auto simp: lookup_bulkload_Some sorted_iff_nth_Suc last_conv_nth + list_all_iff in_set_conv_nth Ball_def Bex_def image_iff split: nat.splits) + +lift_definition trace_mapping_nth :: "'a trace_mapping \ nat \ ('a set \ nat)" is + "\(n, m, t) i. if i < n then the (Mapping.lookup t i) else ({}, (i - n) + m)" . + +lift_definition Trace_Mapping :: "'a trace_mapping \ 'a Trace.trace" is + "\(n, m, t). map (the \ Mapping.lookup t) [0..n. ({} :: 'a set, n + m)) nats" +proof (goal_cases) + case (1 prod) + obtain n m t where prod_def: "prod = (n, m, t)" + by (cases prod) auto + have props: "Mapping.keys t = {.. (the \ Mapping.lookup t)) [0.. True | Suc n' \ (case Mapping.lookup t n' of Some (X', t') \ t' \ m | None \ False))" + "(\n' < n. case Mapping.lookup t n' of Some (X', t') \ finite X' | None \ False)" + using 1 by (auto simp add: prod_def) + have aux: "x \ set (map (the \ Mapping.lookup t) [0.. snd x \ m" for x + using props(2,3) less_Suc_eq_le + by (fastforce simp: sorted_iff_nth_mono split: nat.splits option.splits) + have aux2: "x \ set (map (the \ Mapping.lookup t) [0.. finite (fst x)" for x + using props(1,4) + by (auto split: nat.splits option.splits) + show ?case + unfolding prod_def prod.case + by (rule extend_is_stream[where ?m=m]) (use props aux aux2 in \auto simp: prod_def\) +qed + +code_datatype Trace_Mapping + +definition "trace_of_list xs = Trace_Mapping (trace_mapping_of_list xs)" + +lemma \_rbt_code[code]: "\ (Trace_Mapping t) i = fst (trace_mapping_nth t i)" + by transfer (auto split: prod.splits) + +lemma \_rbt_code[code]: "\ (Trace_Mapping t) i = snd (trace_mapping_nth t i)" + by transfer (auto split: prod.splits) + +lemma trace_mapping_of_list_sound: "sorted (map snd xs) \ (\x \ set xs. finite (fst x)) \ i < length xs \ + xs ! i = (\ (trace_of_list xs) i, \ (trace_of_list xs) i)" + unfolding trace_of_list_def + by transfer (auto simp: lookup_bulkload_Some) + +subsection \Auxiliary results\ + +definition "sum_proofs f xs = sum_list (map f xs)" + +lemma sum_proofs_empty[simp]: "sum_proofs f [] = 0" + by (auto simp: sum_proofs_def) + +lemma sum_proofs_fundef_cong[fundef_cong]: "(\x. x \ set xs \ f x = f' x) \ + sum_proofs f xs = sum_proofs f' xs" + by (induction xs) (auto simp: sum_proofs_def) + +lemma sum_proofs_Cons: + fixes f :: "'a \ nat" + shows "sum_proofs f (p # qs) = f p + sum_proofs f qs" + by (auto simp: sum_proofs_def split: list.splits) + +lemma sum_proofs_app: + fixes f :: "'a \ nat" + shows "sum_proofs f (qs @ [p]) = f p + sum_proofs f qs" + by (auto simp: sum_proofs_def split: list.splits) + +context + fixes w :: "'n \ nat" +begin + +function (sequential) s_pred :: "('n, 'd) sproof \ nat" + and v_pred :: "('n, 'd) vproof \ nat" where + "s_pred (STT _) = 1" +| "s_pred (SEq_Const _ _ _) = 1" +| "s_pred (SPred _ r _) = w r" +| "s_pred (SNeg vp) = (v_pred vp) + 1" +| "s_pred (SOrL sp1) = (s_pred sp1) + 1" +| "s_pred (SOrR sp2) = (s_pred sp2) + 1" +| "s_pred (SAnd sp1 sp2) = (s_pred sp1) + (s_pred sp2) + 1" +| "s_pred (SImpL vp1) = (v_pred vp1) + 1" +| "s_pred (SImpR sp2) = (s_pred sp2) + 1" +| "s_pred (SIffSS sp1 sp2) = (s_pred sp1) + (s_pred sp2) + 1" +| "s_pred (SIffVV vp1 vp2) = (v_pred vp1) + (v_pred vp2) + 1" +| "s_pred (SExists _ _ sp) = (s_pred sp) + 1" +| "s_pred (SForall _ part) = (sum_proofs s_pred (vals part)) + 1" +| "s_pred (SPrev sp) = (s_pred sp) + 1" +| "s_pred (SNext sp) = (s_pred sp) + 1" +| "s_pred (SOnce _ sp) = (s_pred sp) + 1" +| "s_pred (SEventually _ sp) = (s_pred sp) + 1" +| "s_pred (SHistorically _ _ sps) = (sum_proofs s_pred sps) + 1" +| "s_pred (SHistoricallyOut _) = 1" +| "s_pred (SAlways _ _ sps) = (sum_proofs s_pred sps) + 1" +| "s_pred (SSince sp2 sp1s) = (sum_proofs s_pred (sp2 # sp1s)) + 1" +| "s_pred (SUntil sp1s sp2) = (sum_proofs s_pred (sp1s @ [sp2])) + 1" +| "v_pred (VFF _ ) = 1" +| "v_pred (VEq_Const _ _ _) = 1" +| "v_pred (VPred _ r _) = w r" +| "v_pred (VNeg sp) = (s_pred sp) + 1" +| "v_pred (VOr vp1 vp2) = ((v_pred vp1) + (v_pred vp2)) + 1" +| "v_pred (VAndL vp1) = (v_pred vp1) + 1" +| "v_pred (VAndR vp2) = (v_pred vp2) + 1" +| "v_pred (VImp sp1 vp2) = ((s_pred sp1) + (v_pred vp2)) + 1" +| "v_pred (VIffSV sp1 vp2) = ((s_pred sp1) + (v_pred vp2)) + 1" +| "v_pred (VIffVS vp1 sp2) = ((v_pred vp1) + (s_pred sp2)) + 1" +| "v_pred (VExists _ part) = (sum_proofs v_pred (vals part)) + 1" +| "v_pred (VForall _ _ vp) = (v_pred vp) + 1" +| "v_pred (VPrev vp) = (v_pred vp) + 1" +| "v_pred (VPrevZ) = 1" +| "v_pred (VPrevOutL _) = 1" +| "v_pred (VPrevOutR _) = 1" +| "v_pred (VNext vp) = (v_pred vp) + 1" +| "v_pred (VNextOutL _) = 1" +| "v_pred (VNextOutR _) = 1" +| "v_pred (VOnceOut _) = 1" +| "v_pred (VOnce _ _ vps) = (sum_proofs v_pred vps) + 1" +| "v_pred (VEventually _ _ vps) = (sum_proofs v_pred vps) + 1" +| "v_pred (VHistorically _ vp) = (v_pred vp) + 1" +| "v_pred (VAlways _ vp) = (v_pred vp) + 1" +| "v_pred (VSinceOut _) = 1" +| "v_pred (VSince _ vp1 vp2s) = (sum_proofs v_pred (vp1 # vp2s)) + 1" +| "v_pred (VSinceInf _ _ vp2s) = (sum_proofs v_pred vp2s) + 1" +| "v_pred (VUntil _ vp2s vp1) = (sum_proofs v_pred (vp2s @ [vp1])) + 1" +| "v_pred (VUntilInf _ _ vp2s) = (sum_proofs v_pred vp2s) + 1" + by pat_completeness auto +termination + by (relation "measure (case_sum size size)") + (auto simp add: termination_simp) + +definition p_pred :: "('n, 'd) proof \ nat" where + "p_pred = case_sum s_pred v_pred" + +end + +subsection \\<^const>\v_check_exec\ setup\ + +lemma ETP_minus_le_iff: "ETP \ (\ \ i - n) \ j \ \ \ i j \ n" + by (simp add: add.commute i_ETP_tau le_diff_conv) + +lemma ETP_minus_gt_iff: "j < ETP \ (\ \ i - n) \ \ \ i j > n" + by (meson ETP_minus_le_iff leD le_less_linear) + +lemma nat_le_iff_less: + fixes n :: nat + shows "(j \ n) \ (j = 0 \ j - 1 < n)" + by auto + +lemma ETP_minus_eq_iff: "j = ETP \ (\ \ i - n) \ ((j = 0 \ n < \ \ i (j - 1)) \ \ \ i j \ n)" + unfolding eq_iff[of j "ETP \ (\ \ i - n)"] nat_le_iff_less[of j] ETP_minus_le_iff ETP_minus_gt_iff + by auto + +lemma LTP_minus_ge_iff: "\ \ 0 + n \ \ \ i \ j \ LTP \ (\ \ i - n) \ + (case n of 0 \ \ \ j i = 0 | _ \ j \ i \ \ \ i j \ n)" + using \_mono[of i j \] + by (fastforce simp add: i_LTP_tau le_diff_conv2 Suc_le_eq split: nat.splits) + +lemma LTP_plus_ge_iff: "j \ LTP \ (\ \ i + n) \ \ \ j i \ n" + by (simp add: add.commute i_LTP_tau le_diff_conv trans_le_add2) + +lemma LTP_minus_lt_if: + assumes "j \ i" "\ \ 0 + n \ \ \ i" "\ \ i j < n" + shows "LTP \ (\ \ i - n) < j" +proof - + have not_in: "k \ {ia. \ \ ia \ \ \ i - n}" if "j \ k" for k + using assms \_mono[OF that, of \] + by auto + then have "{ia. \ \ ia \ \ \ i - n} \ {0.. \ ia \ \ \ i - n}" + using subset_eq_atLeast0_lessThan_finite + by blast + moreover have "0 \ {ia. \ \ ia \ \ \ i - n}" + using assms(2) + by auto + ultimately show ?thesis + unfolding LTP_def + by (metis Max_in not_in empty_iff not_le_imp_less) +qed + +lemma LTP_minus_lt_iff: + assumes "\ \ 0 + n \ \ \ i" + shows "LTP \ (\ \ i - n) < j \ (if \ j \ i \ n = 0 then \ \ j i > 0 else \ \ i j < n)" +proof (cases "j \ i") + case True + then show ?thesis + using assms i_le_LTPi_minus[of \ n i] LTP_minus_lt_if[of j i \ n] + by (cases n) + (auto simp add: i_LTP_tau linorder_not_less Suc_le_eq dest!: tau_LTP_k[rotated]) +next + case False + have delta: "\ \ i j = 0" + using False + by auto + show ?thesis + proof (cases n) + case 0 + then show ?thesis + using False assms + by (metis add.right_neutral diff_is_0_eq diff_zero i_LTP_tau linorder_not_less) + next + case (Suc n') + then show ?thesis + using False assms + by (cases i) + (auto simp: Suc_le_eq not_le elim!: order.strict_trans[rotated] intro!: i_le_LTPi_minus) + qed +qed + +lemma LTP_minus_eq_iff: + assumes "\ \ 0 + n \ \ \ i" + shows "j = LTP \ (\ \ i - n) \ + (case n of 0 \ i \ j \ \ \ j i = 0 \ \ \ (Suc j) j > 0 + | _ \ j \ i \ n \ \ \ i j \ \ \ i (Suc j) < n)" +proof (cases n) + case 0 + show ?thesis + using assms 0 i_LTP_tau[of \ "\ \ i" "LTP \ (\ \ i)"] + i_LTP_tau[of \ "\ \ i" "Suc (LTP \ (\ \ i))"] i_LTP_tau[of \ "\ \ i" "j"] + less_\D[of \ "(LTP \ (\ \ i))" "Suc j"] + by (auto simp: i_le_LTPi not_le elim!: antisym dest!: + order_antisym_conv[of "\ \ i" "\ \ j", THEN iffD1, rotated] + order_antisym_conv[of "\ \ i" "\ \ (LTP \ (\ \ i))", THEN iffD1, rotated]) +next + case (Suc n') + show ?thesis + using assms + by (simp add: Suc eq_iff[of j "LTP \ (\ \ i - Suc n')"] less_Suc_eq_le[of "LTP \ (\ \ i - Suc n')" j, symmetric] LTP_minus_ge_iff LTP_minus_lt_iff) +qed + +lemma LTP_plus_eq_iff: + shows "j = LTP \ (\ \ i + n) \ (\ \ j i \ n \ \ \ (Suc j) i > n)" + unfolding eq_iff[of j "LTP \ (\ \ i + n)"] + by (meson LTP_plus_ge_iff linorder_not_less not_less_eq_eq) + +lemma LTP_p_def: "\ \ 0 + left I \ \ \ i \ LTP_p \ i I = (case left I of 0 \ i | _ \ LTP \ (\ \ i - left I))" + using i_le_LTPi by (auto simp: min_def i_LTP_tau split: nat.splits) + +definition "check_upt_LTP_p \ I li xs i \ (case xs of [] \ + (case left I of 0 \ i < li | Suc n \ + (if \ li \ i \ left I = 0 then 0 < \ \ li i else \ \ i li < left I)) + | _ \ xs = [li..
  • + (case left I of 0 \ li + length xs - 1 = i | Suc n \ + (li + length xs - 1 \ i \ left I \ \ \ i (li + length xs - 1) \ \ \ i (li + length xs) < left I)))" + +lemma check_upt_l_cong: + assumes "\j. j \ max i li \ \ \ j = \ \' j" + shows "check_upt_LTP_p \ I li xs i = check_upt_LTP_p \' I li xs i" +proof - + have "li + length ys \ i \ Suc n \ \ \' i (li + length ys) \ + (Suc (li + length ys)) \ i" for ys :: "nat list" and n + by (cases "li + length ys = i") auto + then show ?thesis + using assms + by (fastforce simp: check_upt_LTP_p_def Let_def simp del: upt.simps split: list.splits nat.splits) +qed + +lemma check_upt_LTP_p_eq: + assumes "\ \ 0 + left I \ \ \ i" + shows "xs = [li.. i I)] \ check_upt_LTP_p \ I li xs i" +proof - + have "li + length xs = Suc (LTP_p \ i I) \ li + length xs - 1 = LTP_p \ i I" if "xs \ []" + using that + by (cases xs) auto + then have "xs = [li.. i I)] \ (xs = [] \ LTP_p \ i I < li) \ + (xs \ [] \ xs = [li..
  • li + length xs - 1 = LTP_p \ i I)" + by auto + moreover have "\ \ (xs = [] \ + (case left I of 0 \ i < li | Suc n \ + (if \ li \ i \ left I = 0 then 0 < \ \ li i else \ \ i li < left I))) \ + (xs \ [] \ xs = [li..
  • + (case left I of 0 \ li + length xs - 1 = i | Suc n \ + (case left I of 0 \ i \ li + length xs - 1 \ + \ \ (li + length xs - 1) i = 0 \ 0 < \ \ (Suc (li + length xs - 1)) (li + length xs - 1) + | Suc n \ li + length xs - 1 \ i \ + left I \ \ \ i (li + length xs - 1) \ \ \ i (Suc (li + length xs - 1)) < left I)))" + using LTP_p_def[OF assms, symmetric] + unfolding LTP_minus_lt_iff[OF assms, symmetric] + unfolding LTP_minus_eq_iff[OF assms, symmetric] + by (auto split: nat.splits) + moreover have "\ \ (case xs of [] \ + (case left I of 0 \ i < li | Suc n \ + (if \ li \ i \ left I = 0 then 0 < \ \ li i else \ \ i li < left I)) + | _ \ xs = [li..
  • + (case left I of 0 \ li + length xs - 1 = i | Suc n \ + (li + length xs - 1 \ i \ left I \ \ \ i (li + length xs - 1) \ \ \ i (li + length xs) < left I)))" + by (auto split: nat.splits list.splits) + ultimately show ?thesis + unfolding check_upt_LTP_p_def + by simp +qed + +lemma v_check_exec_Once_code[code]: "v_check_exec \ vs (Formula.Once I \) vp = (case vp of + VOnce i li vps \ + (case right I of \ \ li = 0 | enat b \ ((li = 0 \ b < \ \ i (li - 1)) \ \ \ i li \ b)) + \ \ \ 0 + left I \ \ \ i + \ check_upt_LTP_p \ I li (map v_at vps) i \ Ball (set vps) (v_check_exec \ vs \) + | VOnceOut i \ \ \ i < \ \ 0 + left I + | _ \ False)" + by (auto simp: Let_def check_upt_LTP_p_eq ETP_minus_le_iff ETP_minus_eq_iff split: vproof.splits enat.splits simp del: upt_Suc) + +lemma s_check_exec_Historically_code[code]: "s_check_exec \ vs (Formula.Historically I \) vp = (case vp of + SHistorically i li vps \ + (case right I of \ \ li = 0 | enat b \ ((li = 0 \ b < \ \ i (li - 1)) \ \ \ i li \ b)) + \ \ \ 0 + left I \ \ \ i + \ check_upt_LTP_p \ I li (map s_at vps) i \ Ball (set vps) (s_check_exec \ vs \) + | SHistoricallyOut i \ \ \ i < \ \ 0 + left I + | _ \ False)" + by (auto simp: Let_def check_upt_LTP_p_eq ETP_minus_le_iff ETP_minus_eq_iff split: sproof.splits enat.splits simp del: upt_Suc) + +lemma v_check_exec_Since_code[code]: "v_check_exec \ vs (Formula.Since \ I \) vp = (case vp of + VSince i vp1 vp2s \ + let j = v_at vp1 in + (case right I of \ \ True | enat b \ \ \ i j \ b) \ j \ i + \ \ \ 0 + left I \ \ \ i + \ check_upt_LTP_p \ I j (map v_at vp2s) i + \ v_check_exec \ vs \ vp1 \ Ball (set vp2s) (v_check_exec \ vs \) + | VSinceInf i li vp2s \ + (case right I of \ \ li = 0 | enat b \ ((li = 0 \ b < \ \ i (li - 1)) \ \ \ i li \ b)) \ + \ \ 0 + left I \ \ \ i \ + check_upt_LTP_p \ I li (map v_at vp2s) i \ Ball (set vp2s) (v_check_exec \ vs \) + | VSinceOut i \ \ \ i < \ \ 0 + left I + | _ \ False)" + by (auto simp: Let_def check_upt_LTP_p_eq ETP_minus_le_iff ETP_minus_eq_iff split: vproof.splits enat.splits simp del: upt_Suc) + +lemma ETP_f_le_iff: "max i (ETP \ (\ \ i + a)) \ j \ i \ j \ \ \ j i \ a" + by (metis add.commute max.bounded_iff \_mono i_ETP_tau le_diff_conv2) + +lemma ETP_f_ge_iff: "j \ max i (ETP \ (\ \ i + n)) \ (case n of 0 \ j \ i + | Suc n' \ (case j of 0 \ True | Suc j' \ \ \ j' i < n))" +proof (cases n) + case 0 + then show ?thesis + by (auto simp: max_def) (metis i_ge_etpi verit_la_disequality) +next + case (Suc n') + have max: "max i (ETP \ (\ \ i + n)) = ETP \ (\ \ i + n)" + by (auto simp: max_def Suc) + (metis Groups.ab_semigroup_add_class.add.commute ETP_ge less_or_eq_imp_le plus_1_eq_Suc) + have "j \ max i (ETP \ (\ \ i + n)) \ (\ia. \ \ i + n \ \ \ ia \ j \ ia)" + unfolding max + unfolding ETP_def + by (meson LeastI_ex Least_le order.trans ex_le_\) + moreover have "\ \ (case j of 0 \ True | Suc j' \ \\ \ i + n \ \ \ j')" + by (auto split: nat.splits) (meson i_ETP_tau le_trans not_less_eq_eq) + moreover have "\ \ (case j of 0 \ True | Suc j' \ \ \ j' i < n)" + by (auto simp: Suc split: nat.splits) + ultimately show ?thesis + by (auto simp: Suc) +qed + +definition "check_upt_ETP_f \ I i xs hi \ (let j = Suc hi - length xs in + (case xs of [] \ (case left I of 0 \ Suc hi \ i | Suc n' \ \ \ hi i < left I) + | _ \ (xs = [j.. + (case left I of 0 \ j \ i | Suc n' \ + (case j of 0 \ True | Suc j' \ \ \ j' i < left I)) \ + i \ j \ left I \ \ \ j i)))" + +lemma check_upt_lu_cong: + assumes "\j. min i hi \ j \ j \ max i hi \ \ \ j = \ \' j" + shows "check_upt_ETP_f \ I i xs hi = check_upt_ETP_f \' I i xs hi" + using assms + unfolding check_upt_ETP_f_def + by (auto simp add: Let_def le_Suc_eq split: list.splits nat.splits) + +lemma check_upt_ETP_f_eq: "xs = [ETP_f \ i I.. check_upt_ETP_f \ I i xs hi" +proof - + have F1: "(case left I of 0 \ Suc hi \ i | Suc n' \ \ \ hi i < left I) = + (Suc hi \ ETP_f \ i I)" + unfolding ETP_f_ge_iff[where ?j="Suc hi" and ?n="left I"] + by (auto split: nat.splits) + have "xs = [ETP_f \ i I.. (let j = Suc hi - length xs in + (xs = [] \ (case left I of 0 \ Suc hi \ i | Suc n' \ \ \ hi i < left I)) \ + (xs \ [] \ xs = [j.. + (case left I of 0 \ j \ i | Suc n' \ + (case j of 0 \ True | Suc j' \ \ \ j' i < left I)) \ + i \ j \ left I \ \ \ j i))" + unfolding F1 + unfolding Let_def + unfolding ETP_f_ge_iff[where ?n="left I", symmetric] + unfolding ETP_f_le_iff[symmetric] + unfolding eq_iff[of _ "ETP_f \ i I", symmetric] + by auto + moreover have "\ \ (let j = Suc hi - length xs in + (case xs of [] \ (case left I of 0 \ Suc hi \ i | Suc n' \ \ \ hi i < left I) + | _ \ (xs = [j.. + (case left I of 0 \ j \ i | Suc n' \ + (case j of 0 \ True | Suc j' \ \ \ j' i < left I)) \ + i \ j \ left I \ \ \ j i)))" + by (auto simp: Let_def split: list.splits) + finally show ?thesis + unfolding check_upt_ETP_f_def . +qed + +lemma v_check_exec_Eventually_code[code]: "v_check_exec \ vs (Formula.Eventually I \) vp = (case vp of + VEventually i hi vps \ + (case right I of \ \ False | enat b \ (\ \ hi i \ b \ b < \ \ (Suc hi) i)) \ + check_upt_ETP_f \ I i (map v_at vps) hi \ Ball (set vps) (v_check_exec \ vs \) + | _ \ False)" + by (auto simp: Let_def LTP_plus_ge_iff LTP_plus_eq_iff check_upt_ETP_f_eq simp del: upt_Suc + split: vproof.splits enat.splits) + +lemma s_check_exec_Always_code[code]: "s_check_exec \ vs (Formula.Always I \) sp = (case sp of + SAlways i hi sps \ + (case right I of \ \ False | enat b \ (\ \ hi i \ b \ b < \ \ (Suc hi) i)) + \ check_upt_ETP_f \ I i (map s_at sps) hi \ Ball (set sps) (s_check_exec \ vs \) + | _ \ False)" + by (auto simp: Let_def LTP_plus_ge_iff LTP_plus_eq_iff check_upt_ETP_f_eq simp del: upt_Suc + split: sproof.splits enat.splits) + +lemma v_check_exec_Until_code[code]: "v_check_exec \ vs (Formula.Until \ I \) vp = (case vp of + VUntil i vp2s vp1 \ + let j = v_at vp1 in + (case right I of \ \ True | enat b \ j < LTP_f \ i b) + \ i \ j \ check_upt_ETP_f \ I i (map v_at vp2s) j + \ v_check_exec \ vs \ vp1 \ Ball (set vp2s) (v_check_exec \ vs \) + | VUntilInf i hi vp2s \ + (case right I of \ \ False | enat b \ (\ \ hi i \ b \ b < \ \ (Suc hi) i)) \ + check_upt_ETP_f \ I i (map v_at vp2s) hi \ Ball (set vp2s) (v_check_exec \ vs \) + | _ \ False)" + by (auto simp: Let_def LTP_plus_ge_iff LTP_plus_eq_iff check_upt_ETP_f_eq simp del: upt_Suc + split: vproof.splits enat.splits) + +subsection\ETP/LTP setup\ + +lemma ETP_aux: "\ t \ \ \ i \ i \ (LEAST i. t \ \ \ i)" + by (meson LeastI_ex \_mono ex_le_\ nat_le_linear order_trans) + +function ETP_rec where + "ETP_rec \ t i = (if \ \ i \ t then i else ETP_rec \ t (i + 1))" + by pat_completeness auto +termination + using ETP_aux + by (relation "measure (\(\, t, i). Suc (ETP \ t) - i)") + (fastforce simp: ETP_def)+ + +lemma ETP_rec_sound: "ETP_rec \ t j = (LEAST i. i \ j \ t \ \ \ i)" +proof (induction \ t j rule: ETP_rec.induct) + case (1 \ t i) + show ?case + proof (cases "\ \ i \ t") + case True + then show ?thesis + by simp (metis (no_types, lifting) Least_equality order_refl) + next + case False + then show ?thesis + using 1[OF False] + by (simp add: ETP_rec.simps[of _ _ i] del: ETP_rec.simps) + (metis Suc_leD le_antisym not_less_eq_eq) + qed +qed + +lemma ETP_code[code]: "ETP \ t = ETP_rec \ t 0" + using ETP_rec_sound[of \ t 0] + by (auto simp: ETP_def) + +lemma LTP_aux: + assumes "\ \ (Suc i) \ t" + shows "i \ Max {i. \ \ i \ t}" +proof - + have "finite {i. \ \ i \ t}" + by (smt (verit, del_insts) \_mono finite_nat_set_iff_bounded_le i_LTP_tau le0 le_trans mem_Collect_eq) + moreover have "i \ {i. \ \ i \ t}" + using le_trans[OF \_mono[of i "Suc i" \] assms] + by auto + ultimately show ?thesis + by (rule Max_ge) +qed + +function (sequential) LTP_rec where + "LTP_rec \ t i = (if \ \ (Suc i) \ t then LTP_rec \ t (i + 1) else i)" + by pat_completeness auto +termination + using LTP_aux + by (relation "measure (\(\, t, i). Suc (LTP \ t) - i)") (fastforce simp: LTP_def)+ + +lemma LTP_rec_sound: "LTP_rec \ t j = Max ({i. i \ j \ (\ \ i) \ t} \ {j})" +proof (induction \ t j rule: LTP_rec.induct) + case (1 \ t j) + have fin: "finite {i. j \ i \ \ \ i \ t}" + by (smt (verit, del_insts) \_mono finite_nat_set_iff_bounded_le i_LTP_tau le0 le_trans + mem_Collect_eq) + show ?case + proof (cases "\ \ (Suc j) \ t") + case True + have diffI: "{i. Suc j \ i \ \ \ i \ t} = {i. j \ i \ \ \ i \ t} - {j}" + by auto + show ?thesis + using 1[OF True] True fin + by (auto simp del: LTP_rec.simps simp add: LTP_rec.simps[of _ _ j] diffI intro: max_aux) + next + case False + then show ?thesis + using fin + by (auto simp: not_le intro!: Max_insert2[symmetric] + dest!: order.strict_trans1 less_\D) + qed +qed + +lemma LTP_code[code]: "LTP \ t = (if t < \ \ 0 + then Code.abort (STR ''LTP: undefined'') (\_. LTP \ t) + else LTP_rec \ t 0)" + using LTP_rec_sound[of \ t 0] + by (auto simp: LTP_def insert_absorb simp del: LTP_rec.simps) + +lemma map_part_code[code]: "Rep_part (map_part f xs) = map (map_prod id f) (Rep_part xs)" + using Rep_part[of xs] + by (auto simp: map_part_def intro!: Abs_part_inverse) + +lemma coset_subset_set_code[code]: + "(List.coset (xs :: _ :: universe list) \ set ys) = (case universe of None \ False + | Some zs \ \z \ set zs. z \ set xs \ z \ set ys)" + using finite_compl finite_subset + by (auto split: option.splits dest!: infinite finite) + +lemma is_empty_coset[code]: "Set.is_empty (List.coset (xs :: _ :: universe list)) = + (case universe of None \ False + | Some zs \ \z \ set zs. z \ set xs)" + using coset_subset_set_code[of xs] by (auto simp: Set.is_empty_def split: option.splits dest: infinite finite) + +subsection \Exported functions\ + +type_synonym name = string8 + +declare Formula.future_bounded.simps[code] + +definition collect_paths :: "('n, 'd :: {default, linorder}) trace \ ('n, 'd) formula \ ('n, 'd) expl \ 'd set list set option" where + "collect_paths \ \ e = (if (distinct_paths e \ check_all_aux \ (\_. UNIV) \ e) then None else Some (collect_paths_aux {[]} \ (\_. UNIV) \ e))" + +definition check :: "(name, event_data) trace \ (name, event_data) formula \ (name, event_data) expl \ bool" where + "check = check_all" + +definition collect_paths_specialized :: "(name, event_data) trace \ (name, event_data) formula \ (name, event_data) expl \ event_data set list set option" where + "collect_paths_specialized = collect_paths" + +definition trace_of_list_specialized :: "((name \ event_data list) set \ nat) list \ (name, event_data) trace" where + "trace_of_list_specialized xs = trace_of_list xs" + +definition specialized_set :: "(name \ event_data list) list \ (name \ event_data list) set" where + "specialized_set = set" + +definition ed_set :: "event_data list \ event_data set" where + "ed_set = set" + +definition sum_nat :: "nat \ nat \ nat" where + "sum_nat m n = m + n" + +definition sub_nat :: "nat \ nat \ nat" where + "sub_nat m n = m - n" + +lift_definition abs_part :: "(event_data set \ 'a) list \ (event_data, 'a) part" is + "\xs. + let Ds = map fst xs in + if {} \ set Ds + \ (\D \ set Ds. \E \ set Ds. D \ E \ D \ E \ {}) + \ \ distinct Ds + \ (\D \ set Ds. D) \ UNIV then [(UNIV, undefined)] else xs" + by (auto simp: partition_on_def disjoint_def) + +export_code interval enat nat_of_integer integer_of_nat + STT Formula.TT Inl EInt Formula.Var Leaf set part_hd sum_nat sub_nat subsvals + check trace_of_list_specialized specialized_set ed_set abs_part + collect_paths_specialized + in OCaml module_name Checker file_prefix "checker" + +(*<*) +end +(*>*) diff --git a/thys/MFOTL_Checker/Event_Data.thy b/thys/MFOTL_Checker/Event_Data.thy new file mode 100644 --- /dev/null +++ b/thys/MFOTL_Checker/Event_Data.thy @@ -0,0 +1,173 @@ +(*<*) +theory Event_Data + imports + "HOL-Library.Char_ord" +begin +(*>*) + +section \Type of Events\ + +subsection \Code Adaptation for 8-bit strings\ + +typedef string8 = "UNIV :: char list set" .. + +setup_lifting type_definition_string8 + +lift_definition empty_string :: string8 is "[]" . +lift_definition string8_literal :: "String.literal \ string8" is String.explode . +lift_definition literal_string8:: "string8 \ String.literal" is String.Abs_literal . +declare [[coercion string8_literal]] + +instantiation string8 :: "{equal, linorder}" +begin + +lift_definition equal_string8 :: "string8 \ string8 \ bool" is HOL.equal . +lift_definition less_eq_string8 :: "string8 \ string8 \ bool" is ord_class.lexordp_eq . +lift_definition less_string8 :: "string8 \ string8 \ bool" is ord_class.lexordp . + +instance by intro_classes + (transfer; auto simp: equal_eq lexordp_conv_lexordp_eq lexordp_eq_linear + intro: lexordp_eq_refl lexordp_eq_trans lexordp_eq_antisym)+ + +end + +lifting_forget string8.lifting + +declare [[code drop: literal_string8 string8_literal "HOL.equal :: string8 \ _" + "(\) :: string8 \ _" "(<) :: string8 \ _" + "Code_Evaluation.term_of :: string8 \ _"]] + +code_printing + type_constructor string8 \ (OCaml) "string" + | constant "HOL.equal :: string8 \ string8 \ bool" \ (OCaml) "Stdlib.(=)" + | constant "(\) :: string8 \ string8 \ bool" \ (OCaml) "Stdlib.(<=)" + | constant "(<) :: string8 \ string8 \ bool" \ (OCaml) "Stdlib.(<)" + | constant "empty_string :: string8" \ (OCaml) "\"\"" + | constant "string8_literal :: String.literal \ string8" \ (OCaml) "id" + | constant "literal_string8 :: string8 \ String.literal" \ (OCaml) "id" + +ML \structure String8 = struct fun to_term x = @{term Abs_string8} $ HOLogic.mk_string x; end;\ + +code_printing + type_constructor string8 \ (Eval) "string" + | constant "string8_literal :: String.literal \ string8" \ (Eval) "_" + | constant "HOL.equal :: string8 \ string8 \ bool" \ (Eval) infixl 6 "=" + | constant "(\) :: string8 \ string8 \ bool" \ (Eval) infixl 6 "<=" + | constant "(<) :: string8 \ string8 \ bool" \ (Eval) infixl 6 "<" + | constant "empty_string :: string8" \ (Eval) "\"\"" + | constant "Code_Evaluation.term_of :: string8 \ term" \ (Eval) "String8.to'_term" + +ML \structure String8 =struct fun to_term x = @{term Abs_string8} $ HOLogic.mk_string x; end;\ + +code_printing + type_constructor string8 \ (Eval) "string" + | constant "string8_literal :: String.literal \ string8" \ (Eval) "_" + | constant "HOL.equal :: string8 \ string8 \ bool" \ (Eval) infixl 6 "=" + | constant "(\) :: string8 \ string8 \ bool" \ (Eval) infixl 6 "<=" + | constant "(<) :: string8 \ string8 \ bool" \ (Eval) infixl 6 "<" + | constant "Code_Evaluation.term_of :: string8 \ term" \ (Eval) "String8.to'_term" + +subsection \Event Parameters\ + +definition div_to_zero :: "integer \ integer \ integer" where + "div_to_zero x y = (let z = fst (Code_Numeral.divmod_abs x y) in + if (x < 0) \ (y < 0) then - z else z)" + +definition mod_to_zero :: "integer \ integer \ integer" where + "mod_to_zero x y = (let z = snd (Code_Numeral.divmod_abs x y) in + if x < 0 then - z else z)" + +lemma "b \ 0 \ div_to_zero a b * b + mod_to_zero a b = a" + unfolding div_to_zero_def mod_to_zero_def Let_def + by (auto simp: minus_mod_eq_mult_div[symmetric] div_minus_right mod_minus_right ac_simps) + +datatype event_data = EInt integer | EString string8 + +instantiation event_data :: "{ord, plus, minus, uminus, times, divide, modulo}" +begin + +fun less_eq_event_data where + "EInt x \ EInt y \ x \ y" +| "EString x \ EString y \ x \ y" +| "EInt _ \ EString _ \ True" +| "(_ :: event_data) \ _ \ False" + +definition less_event_data :: "event_data \ event_data \ bool" where + "less_event_data x y \ x \ y \ \ y \ x" + +fun plus_event_data where + "EInt x + EInt y = EInt (x + y)" +| "(_::event_data) + _ = undefined" + +fun minus_event_data where + "EInt x - EInt y = EInt (x - y)" +| "(_::event_data) - _ = undefined" + +fun uminus_event_data where + "- EInt x = EInt (- x)" +| "- (_::event_data) = undefined" + +fun times_event_data where + "EInt x * EInt y = EInt (x * y)" +| "(_::event_data) * _ = undefined" + +fun divide_event_data where + "EInt x div EInt y = EInt (div_to_zero x y)" +| "(_::event_data) div _ = undefined" + +fun modulo_event_data where + "EInt x mod EInt y = EInt (mod_to_zero x y)" +| "(_::event_data) mod _ = undefined" + +instance .. + +end + +lemma infinite_UNIV_event_data: + "\finite (UNIV :: event_data set)" +proof - + define f where "f = (\k. EInt k)" + have inj: "inj_on f (UNIV :: integer set)" + unfolding f_def by (meson event_data.inject(1) injI) + show ?thesis using finite_imageD[OF _ inj] + by (meson infinite_UNIV_char_0 infinite_iff_countable_subset top_greatest) +qed + +primrec integer_of_event_data :: "event_data \ integer" where + "integer_of_event_data (EInt _) = undefined" +| "integer_of_event_data (EString _) = undefined" + +instantiation event_data :: default begin + +definition default_event_data :: event_data where "default = EInt 0" + +instance proof qed + +end + +instantiation event_data :: linorder begin +instance +proof (standard, unfold less_event_data_def, goal_cases less refl trans antisym total) + case (refl x) + then show ?case + by (cases x) auto +next + case (trans x y z) + then show ?case + by (cases x; cases y; cases z) auto +next + case (antisym x y) + then show ?case + by (cases x; cases y) auto +next + case (total x y) + then show ?case + by (cases x; cases y) auto +qed simp + +end + + +(*<*) +end +(*>*) diff --git a/thys/MFOTL_Checker/Examples.thy b/thys/MFOTL_Checker/Examples.thy new file mode 100644 --- /dev/null +++ b/thys/MFOTL_Checker/Examples.thy @@ -0,0 +1,104 @@ +(*<*) +theory Examples +imports Monitor Checker_Code +begin +(*>*) + +section \Examples\ + +definition monitor :: "(('n :: linorder \ 'd :: {default, linorder} list) set \ nat) list \ ('n, 'd) formula \ ('n, 'd) expl list" where + "monitor \ \ = map (\i. eval (trace_of_list \) (\p q. size p \ size q) (sorted_list_of_set (fv \)) i \) [0 ..< length \]" +definition check :: "(('n :: linorder \ 'd :: {default, linorder} list) set \ nat) list \ ('n, 'd) formula \ bool" where + "check \ \ = list_all (check_all (trace_of_list \) \) (monitor \ \)" + +subsection \Infinite Domain\ + +definition prefix :: "((string \ string list) set \ nat) list" where + "prefix = + [({(''mgr_S'', [''Mallory'', ''Alice'']), + (''mgr_S'', [''Merlin'', ''Bob'']), + (''mgr_S'', [''Merlin'', ''Charlie''])}, 1307532861::nat), + ({(''approve'', [''Mallory'', ''152''])}, 1307532861), + ({(''approve'', [''Merlin'', ''163'']), + (''publish'', [''Alice'', ''160'']), + (''mgr_F'', [''Merlin'', ''Charlie''])}, 1307955600), + ({(''approve'', [''Merlin'', ''187'']), + (''publish'', [''Bob'', ''163'']), + (''publish'', [''Alice'', ''163'']), + (''publish'', [''Charlie'', ''163'']), + (''publish'', [''Charlie'', ''152''])}, 1308477599)]" + +definition phi :: "(string, string) Formula.formula" where + "phi = Formula.Imp (Formula.Pred ''publish'' [Formula.Var ''a'', Formula.Var ''f'']) + (Formula.Once (init 604800) (Formula.Exists ''m'' (Formula.Since + (Formula.Neg (Formula.Pred ''mgr_F'' [Formula.Var ''m'', Formula.Var ''a''])) all + (Formula.And (Formula.Pred ''mgr_S'' [Formula.Var ''m'', Formula.Var ''a'']) + (Formula.Pred ''approve'' [Formula.Var ''m'', Formula.Var ''f''])))))" + +value "monitor prefix phi" +lemma "check prefix phi" + by eval + +subsection \Finite Domain\ + +datatype Domain = Mallory | Merlin | Martin | Alice | Bob | Charlie | David | Default | R42 | R152 | R160 | R163 | R187 + +definition ord :: "Domain \ nat" where + "ord d = (case d of + Mallory \ 0 + | Merlin \ 1 + | Martin \ 2 + | Alice \ 3 + | Bob \ 4 + | Charlie \ 5 + | David \ 6 + | Default \ 7 + | R42 \ 8 + | R152 \ 9 + | R160 \ 10 + | R163 \ 11 + | R187 \ 12)" + +instantiation Domain :: default begin +definition "default_Domain = Default" +instance .. +end +instantiation Domain :: universe begin +definition "universe_Domain = Some [Mallory, Merlin, Martin, Alice, Bob, Charlie, David, Default, R42, R152, R160, R163, R187]" +instance by standard (auto simp: universe_Domain_def intro: Domain.exhaust) +end +instantiation Domain :: linorder begin +definition "less_eq_Domain d d' = (ord d \ ord d')" +definition "less_Domain d d' = (ord d < ord d')" +instance by standard (auto simp: less_eq_Domain_def less_Domain_def ord_def split: Domain.splits) +end + +definition fprefix :: "((string \ Domain list) set \ nat) list" where + "fprefix = + [({(''mgr_S'', [Mallory, Alice]), + (''mgr_S'', [Merlin, Bob]), + (''mgr_S'', [Merlin, Charlie])}, 1307532861::nat), + ({(''approve'', [Mallory, R152])}, 1307532861), + ({(''approve'', [Merlin, R163]), + (''publish'', [Alice, R160]), + (''mgr_F'', [Merlin, Charlie])}, 1307955600), + ({(''approve'', [Merlin, R187]), + (''publish'', [Bob, R163]), + (''publish'', [Alice, R163]), + (''publish'', [Charlie, R163]), + (''publish'', [Charlie, R152])}, 1308477599)]" + +definition fphi :: "(string, Domain) Formula.formula" where + "fphi = Formula.Imp (Formula.Pred ''publish'' [Formula.Var ''a'', Formula.Var ''f'']) + (Formula.Once (init 604800) (Formula.Exists ''m'' (Formula.Since + (Formula.Neg (Formula.Pred ''mgr_F'' [Formula.Var ''m'', Formula.Var ''a''])) all + (Formula.And (Formula.Pred ''mgr_S'' [Formula.Var ''m'', Formula.Var ''a'']) + (Formula.Pred ''approve'' [Formula.Var ''m'', Formula.Var ''f''])))))" + +value "monitor fprefix fphi" +lemma "check fprefix fphi" + by eval + +(*<*) +end +(*>*) \ No newline at end of file diff --git a/thys/MFOTL_Checker/Formula.thy b/thys/MFOTL_Checker/Formula.thy new file mode 100644 --- /dev/null +++ b/thys/MFOTL_Checker/Formula.thy @@ -0,0 +1,370 @@ +(*<*) +theory Formula + imports Trace +begin +(*>*) + +section \Metric First-Order Temporal Logic\ + +subsection \Syntax\ + +type_synonym ('n, 'a) event = "('n \ 'a list)" +type_synonym ('n, 'a) database = "('n, 'a) event set" +type_synonym ('n, 'a) prefix = "('n \ 'a list) prefix" +type_synonym ('n, 'a) trace = "('n \ 'a list) trace" +type_synonym ('n, 'a) env = "'n \ 'a" +type_synonym ('n, 'a) envset = "'n \ 'a set" + +datatype (fv_trm: 'n, 'a) trm = is_Var: Var 'n ("\<^bold>v") | is_Const: Const 'a ("\<^bold>c") + +lemma in_fv_trm_conv: "x \ fv_trm t \ t = \<^bold>v x" + by (cases t) auto + +datatype ('n, 'a) formula = + TT ("\") +| FF ("\") +| Eq_Const 'n 'a ("_ \<^bold>\ _" [85, 85] 85) +| Pred 'n "('n, 'a) trm list" ("_ \ _" [85, 85] 85) +| Neg "('n, 'a) formula" ("\\<^sub>F _" [82] 82) +| Or "('n, 'a) formula" "('n, 'a) formula" (infixr "\\<^sub>F" 80) +| And "('n, 'a) formula" "('n, 'a) formula" (infixr "\\<^sub>F" 80) +| Imp "('n, 'a) formula" "('n, 'a) formula" (infixr "\\<^sub>F" 79) +| Iff "('n, 'a) formula" "('n, 'a) formula" (infixr "\\<^sub>F" 79) +| Exists "'n" "('n, 'a) formula" ("\\<^sub>F_. _" [70,70] 70) +| Forall "'n" "('n, 'a) formula" ("\\<^sub>F_. _" [70,70] 70) +| Prev \ "('n, 'a) formula" ("\<^bold>Y _ _" [1000, 65] 65) +| Next \ "('n, 'a) formula" ("\<^bold>X _ _" [1000, 65] 65) +| Once \ "('n, 'a) formula" ("\<^bold>P _ _" [1000, 65] 65) +| Historically \ "('n, 'a) formula" ("\<^bold>H _ _" [1000, 65] 65) +| Eventually \ "('n, 'a) formula" ("\<^bold>F _ _" [1000, 65] 65) +| Always \ "('n, 'a) formula" ("\<^bold>G _ _" [1000, 65] 65) +| Since "('n, 'a) formula" \ "('n, 'a) formula" ("_ \<^bold>S _ _" [60,1000,60] 60) +| Until "('n, 'a) formula" \ "('n, 'a) formula" ("_ \<^bold>U _ _" [60,1000,60] 60) + +primrec fv :: "('n, 'a) formula \ 'n set" where + "fv (r \ ts) = \ (fv_trm ` set ts)" +| "fv \ = {}" +| "fv \ = {}" +| "fv (x \<^bold>\ c) = {x}" +| "fv (\\<^sub>F \) = fv \" +| "fv (\ \\<^sub>F \) = fv \ \ fv \" +| "fv (\ \\<^sub>F \) = fv \ \ fv \" +| "fv (\ \\<^sub>F \) = fv \ \ fv \" +| "fv (\ \\<^sub>F \) = fv \ \ fv \" +| "fv (\\<^sub>Fx. \) = fv \ - {x}" +| "fv (\\<^sub>Fx. \) = fv \ - {x}" +| "fv (\<^bold>Y I \) = fv \" +| "fv (\<^bold>X I \) = fv \" +| "fv (\<^bold>P I \) = fv \" +| "fv (\<^bold>H I \) = fv \" +| "fv (\<^bold>F I \) = fv \" +| "fv (\<^bold>G I \) = fv \" +| "fv (\ \<^bold>S I \) = fv \ \ fv \" +| "fv (\ \<^bold>U I \) = fv \ \ fv \" + +primrec "consts" :: "('n, 'a) formula \ 'a set" where + "consts (r \ ts) = {}" \ \terms may also contain constants, + but these only filter out values from the trace and do not introduce + new values of interest (i.e., do not extend the active domain)\ +| "consts \ = {}" +| "consts \ = {}" +| "consts (x \<^bold>\ c) = {c}" +| "consts (\\<^sub>F \) = consts \" +| "consts (\ \\<^sub>F \) = consts \ \ consts \" +| "consts (\ \\<^sub>F \) = consts \ \ consts \" +| "consts (\ \\<^sub>F \) = consts \ \ consts \" +| "consts (\ \\<^sub>F \) = consts \ \ consts \" +| "consts (\\<^sub>Fx. \) = consts \" +| "consts (\\<^sub>Fx. \) = consts \" +| "consts (\<^bold>Y I \) = consts \" +| "consts (\<^bold>X I \) = consts \" +| "consts (\<^bold>P I \) = consts \" +| "consts (\<^bold>H I \) = consts \" +| "consts (\<^bold>F I \) = consts \" +| "consts (\<^bold>G I \) = consts \" +| "consts (\ \<^bold>S I \) = consts \ \ consts \" +| "consts (\ \<^bold>U I \) = consts \ \ consts \" + +lemma finite_fv_trm[simp]: "finite (fv_trm t)" + by (cases t) simp_all + +lemma finite_fv[simp]: "finite (fv \)" + by (induction \) simp_all + +lemma finite_consts[simp]: "finite (consts \)" + by (induction \) simp_all + +definition nfv :: "('n, 'a) formula \ nat" where + "nfv \ = card (fv \)" + +fun future_bounded :: "('n, 'a) formula \ bool" where + "future_bounded \ = True" +| "future_bounded \ = True" +| "future_bounded (_ \ _) = True" +| "future_bounded (_ \<^bold>\ _) = True" +| "future_bounded (\\<^sub>F \) = future_bounded \" +| "future_bounded (\ \\<^sub>F \) = (future_bounded \ \ future_bounded \)" +| "future_bounded (\ \\<^sub>F \) = (future_bounded \ \ future_bounded \)" +| "future_bounded (\ \\<^sub>F \) = (future_bounded \ \ future_bounded \)" +| "future_bounded (\ \\<^sub>F \) = (future_bounded \ \ future_bounded \)" +| "future_bounded (\\<^sub>F_. \) = future_bounded \" +| "future_bounded (\\<^sub>F_. \) = future_bounded \" +| "future_bounded (\<^bold>Y I \) = future_bounded \" +| "future_bounded (\<^bold>X I \) = future_bounded \" +| "future_bounded (\<^bold>P I \) = future_bounded \" +| "future_bounded (\<^bold>H I \) = future_bounded \" +| "future_bounded (\<^bold>F I \) = (future_bounded \ \ right I \ \)" +| "future_bounded (\<^bold>G I \) = (future_bounded \ \ right I \ \)" +| "future_bounded (\ \<^bold>S I \) = (future_bounded \ \ future_bounded \)" +| "future_bounded (\ \<^bold>U I \) = (future_bounded \ \ future_bounded \ \ right I \ \)" + +subsection \Semantics\ + +primrec eval_trm :: "('n, 'a) env \ ('n, 'a) trm \ 'a"("_\_\" [70,89] 89) where + "eval_trm v (\<^bold>v x) = v x" +| "eval_trm v (\<^bold>c x) = x" + +lemma eval_trm_fv_cong: "\x\fv_trm t. v x = v' x \ v\t\ = v'\t\" + by (induction t) simp_all + +definition eval_trms :: "('n, 'a) env \ ('n, 'a) trm list \ 'a list" ("_\<^bold>\_\<^bold>\" [70,89] 89) where + "eval_trms v ts = map (eval_trm v) ts" + +lemma eval_trms_fv_cong: + "\t\set ts. \x\fv_trm t. v x = v' x \ v\<^bold>\ts\<^bold>\ = v'\<^bold>\ts\<^bold>\" + using eval_trm_fv_cong[of _ v v'] + by (auto simp: eval_trms_def) + +(* vs :: "'a envset" is used whenever we define executable functions *) +primrec eval_trm_set :: "('n, 'a) envset \ ('n, 'a) trm \ ('n, 'a) trm \ 'a set"("_\_\" [70,89] 89) where + "eval_trm_set vs (\<^bold>v x) = (\<^bold>v x, vs x)" +| "eval_trm_set vs (\<^bold>c x) = (\<^bold>c x, {x})" + +definition eval_trms_set :: "('n, 'a) envset \ ('n, 'a) trm list \ (('n, 'a) trm \ 'a set) list" ("_\<^bold>\_\<^bold>\" [70,89] 89) + where "eval_trms_set vs ts = map (eval_trm_set vs) ts" + +lemma eval_trms_set_Nil: "vs\<^bold>\[]\<^bold>\ = []" + by (simp add: eval_trms_set_def) + +lemma eval_trms_set_Cons: + "vs\<^bold>\(t # ts)\<^bold>\ = vs\t\ # vs\<^bold>\ts\<^bold>\" + by (simp add: eval_trms_set_def) + +primrec sat :: "('n, 'a) trace \ ('n, 'a) env \ nat \ ('n, 'a) formula \ bool" ("\_, _, _\ \ _" [56, 56, 56, 56] 55) where + "\\, v, i\ \ \ = True" +| "\\, v, i\ \ \ = False" +| "\\, v, i\ \ r \ ts = ((r, v\<^bold>\ts\<^bold>\) \ \ \ i)" +| "\\, v, i\ \ x \<^bold>\ c = (v x = c)" +| "\\, v, i\ \ \\<^sub>F \ = (\ \\, v, i\ \ \)" +| "\\, v, i\ \ \ \\<^sub>F \ = (\\, v, i\ \ \ \ \\, v, i\ \ \)" +| "\\, v, i\ \ \ \\<^sub>F \ = (\\, v, i\ \ \ \ \\, v, i\ \ \)" +| "\\, v, i\ \ \ \\<^sub>F \ = (\\, v, i\ \ \ \ \\, v, i\ \ \)" +| "\\, v, i\ \ \ \\<^sub>F \ = (\\, v, i\ \ \ \ \\, v, i\ \ \)" +| "\\, v, i\ \ \\<^sub>Fx. \ = (\z. \\, v(x := z), i\ \ \)" +| "\\, v, i\ \ \\<^sub>Fx. \ = (\z. \\, v(x := z), i\ \ \)" +| "\\, v, i\ \ \<^bold>Y I \ = (case i of 0 \ False | Suc j \ mem (\ \ i - \ \ j) I \ \\, v, j\ \ \)" +| "\\, v, i\ \ \<^bold>X I \ = (mem (\ \ (Suc i) - \ \ i) I \ \\, v, Suc i\ \ \)" +| "\\, v, i\ \ \<^bold>P I \ = (\j\i. mem (\ \ i - \ \ j) I \ \\, v, j\ \ \)" +| "\\, v, i\ \ \<^bold>H I \ = (\j\i. mem (\ \ i - \ \ j) I \ \\, v, j\ \ \)" +| "\\, v, i\ \ \<^bold>F I \ = (\j\i. mem (\ \ j - \ \ i) I \ \\, v, j\ \ \)" +| "\\, v, i\ \ \<^bold>G I \ = (\j\i. mem (\ \ j - \ \ i) I \ \\, v, j\ \ \)" +| "\\, v, i\ \ \ \<^bold>S I \ = (\j\i. mem (\ \ i - \ \ j) I \ \\, v, j\ \ \ \ (\k\{j<..i}. \\, v, k\ \ \))" +| "\\, v, i\ \ \ \<^bold>U I \ = (\j\i. mem (\ \ j - \ \ i) I \ \\, v, j\ \ \ \ (\k\{i..\, v, k\ \ \))" + +lemma sat_fv_cong: "\x\fv \. v x = v' x \ \\, v, i\ \ \ = \\, v', i\ \ \" +proof (induct \ arbitrary: v v' i) + case (Pred n ts) + thus ?case + by (simp cong: map_cong eval_trms_fv_cong[rule_format, OF Pred[simplified, rule_format]] + split: option.splits) +next + case (Exists t \) + then show ?case unfolding sat.simps + by (intro iff_exI) (simp add: nth_Cons') +next + case (Forall t \) + then show ?case unfolding sat.simps + by (intro iff_allI) (simp add: nth_Cons') +qed (auto 10 0 simp: Let_def split: nat.splits intro!: iff_exI eval_trm_fv_cong) + +lemma sat_Until_rec: "\\, v, i\ \ \ \<^bold>U I \ \ + (mem 0 I \ \\, v, i\ \ \ \ + \ \ (i + 1) \ right I \ \\, v, i\ \ \ \ \\, v, i + 1\ \ \ \<^bold>U (subtract (\ \ (i + 1)) I) \)" + (is "?L \ ?R") +proof (rule iffI; (elim disjE conjE)?) + assume ?L + then obtain j where j: "i \ j" "mem (\ \ j - \ \ i) I" "\\, v, j\ \ \" "\k \ {i ..< j}. \\, v, k\ \ \" + by auto + then show ?R + proof (cases "i = j") + case False + with j(1,2) have "\ \ (i + 1) \ right I" + by (auto elim: order_trans[rotated] simp: diff_le_mono) + moreover from False j(1,4) have "\\, v, i\ \ \" by auto + moreover from False j have "\\, v, i + 1\ \ \ \<^bold>U (subtract (\ \ (i + 1)) I) \" + by (cases "right I") (auto simp: le_diff_conv le_diff_conv2 intro!: exI[of _ j]) + ultimately show ?thesis by blast + qed simp +next + assume \: "\ \ (i + 1) \ right I" and now: "\\, v, i\ \ \" and + "next": "\\, v, i + 1\ \ \ \<^bold>U (subtract (\ \ (i + 1)) I) \" + from "next" obtain j where j: "i + 1 \ j" "mem (\ \ j - \ \ (i + 1)) (subtract (\ \ (i + 1)) I)" + "\\, v, j\ \ \" "\k \ {i + 1 ..< j}. \\, v, k\ \ \" + by auto + from \ j(1,2) have "mem (\ \ j - \ \ i) I" + by (cases "right I") (auto simp: le_diff_conv2) + with now j(1,3,4) show ?L by (auto simp: le_eq_less_or_eq[of i] intro!: exI[of _ j]) +qed auto + +lemma sat_Since_rec: "\\, v, i\ \ \ \<^bold>S I \ \ + mem 0 I \ \\, v, i\ \ \ \ + (i > 0 \ \ \ i \ right I \ \\, v, i\ \ \ \ \\, v, i - 1\ \ \ \<^bold>S (subtract (\ \ i) I) \)" + (is "?L \ ?R") +proof (rule iffI; (elim disjE conjE)?) + assume ?L + then obtain j where j: "j \ i" "mem (\ \ i - \ \ j) I" "\\, v, j\ \ \" "\k \ {j <.. i}. \\, v, k\ \ \" + by auto + then show ?R + proof (cases "i = j") + case False + with j(1) obtain k where [simp]: "i = k + 1" + by (cases i) auto + with j(1,2) False have "\ \ i \ right I" + by (auto elim: order_trans[rotated] simp: diff_le_mono2 le_Suc_eq) + moreover from False j(1,4) have "\\, v, i\ \ \" by auto + moreover from False j have "\\, v, i - 1\ \ \ \<^bold>S (subtract (\ \ i) I) \" + by (cases "right I") (auto simp: le_diff_conv le_diff_conv2 intro!: exI[of _ j]) + ultimately show ?thesis by auto + qed simp +next + assume i: "0 < i" and \: "\ \ i \ right I" and now: "\\, v, i\ \ \" and + "prev": "\\, v, i - 1\ \ \ \<^bold>S (subtract (\ \ i) I) \" + from "prev" obtain j where j: "j \ i - 1" "mem (\ \ (i - 1) - \ \ j) ((subtract (\ \ i) I))" + "\\, v, j\ \ \" "\k \ {j <.. i - 1}. \\, v, k\ \ \" + by auto + from \ i j(1,2) have "mem (\ \ i - \ \ j) I" + by (cases "right I") (auto simp: le_diff_conv2) + with now i j(1,3,4) show ?L by (auto simp: le_Suc_eq gr0_conv_Suc intro!: exI[of _ j]) +qed auto + +lemma sat_Since_0: "\\, v, 0\ \ \ \<^bold>S I \ \ mem 0 I \ \\, v, 0\ \ \" + by auto + +lemma sat_Since_point: "\\, v, i\ \ \ \<^bold>S I \ \ + (\j. j \ i \ mem (\ \ i - \ \ j) I \ \\, v, i\ \ \ \<^bold>S (point (\ \ i - \ \ j)) \ \ P) \ P" + by (auto intro: diff_le_self) + +lemma sat_Since_pointD: "\\, v, i\ \ \ \<^bold>S (point t) \ \ mem t I \ \\, v, i\ \ \ \<^bold>S I \" + by auto + +lemma sat_Once_Since: "\\, v, i\ \ \<^bold>P I \ = \\, v, i\ \ \ \<^bold>S I \" + by auto + +lemma sat_Once_rec: "\\, v, i\ \ \<^bold>P I \ \ + mem 0 I \ \\, v, i\ \ \ \ + (i > 0 \ \ \ i \ right I \ \\, v, i - 1\ \ \<^bold>P (subtract (\ \ i) I) \)" + unfolding sat_Once_Since + by (subst sat_Since_rec) auto + +lemma sat_Historically_Once: "\\, v, i\ \ \<^bold>H I \ = \\, v, i\ \ \\<^sub>F (\<^bold>P I \\<^sub>F \)" + by auto + +lemma sat_Historically_rec: "\\, v, i\ \ \<^bold>H I \ \ + (mem 0 I \ \\, v, i\ \ \) \ + (i > 0 \ \ \ i \ right I \ \\, v, i - 1\ \ \<^bold>H (subtract (\ \ i) I) \)" + unfolding sat_Historically_Once sat.simps(5) + by (subst sat_Once_rec) auto + +lemma sat_Eventually_Until: "\\, v, i\ \ \<^bold>F I \ = \\, v, i\ \ \ \<^bold>U I \" + by auto + +lemma sat_Eventually_rec: "\\, v, i\ \ \<^bold>F I \ \ + mem 0 I \ \\, v, i\ \ \ \ + (\ \ (i + 1) \ right I \ \\, v, i + 1\ \ \<^bold>F (subtract (\ \ (i + 1)) I) \)" + unfolding sat_Eventually_Until + by (subst sat_Until_rec) auto + +lemma sat_Always_Eventually: "\\, v, i\ \ \<^bold>G I \ = \\, v, i\ \ \\<^sub>F (\<^bold>F I \\<^sub>F \)" + by auto + +lemma sat_Always_rec: "\\, v, i\ \ \<^bold>G I \ \ + (mem 0 I \ \\, v, i\ \ \) \ + (\ \ (i + 1) \ right I \ \\, v, i + 1\ \ \<^bold>G (subtract (\ \ (i + 1)) I) \)" + unfolding sat_Always_Eventually sat.simps(5) + by (subst sat_Eventually_rec) auto + +bundle MFOTL_no_notation begin + +text \ For bold font, type ``backslash'' followed by the word ``bold'' \ +no_notation Var ("\<^bold>v") + and Const ("\<^bold>c") + +text \ For subscripts type ``backslash'' followed by ``sub'' \ +no_notation TT ("\") + and FF ("\") + and Pred ("_ \ _" [85, 85] 85) + and Eq_Const ("_ \<^bold>\ _" [85, 85] 85) + and Neg ("\\<^sub>F _" [82] 82) + and And (infixr "\\<^sub>F" 80) + and Or (infixr "\\<^sub>F" 80) + and Imp (infixr "\\<^sub>F" 79) + and Iff (infixr "\\<^sub>F" 79) + and Exists ("\\<^sub>F_. _" [70,70] 70) + and Forall ("\\<^sub>F_. _" [70,70] 70) + and Prev ("\<^bold>Y _ _" [1000, 65] 65) + and Next ("\<^bold>X _ _" [1000, 65] 65) + and Once ("\<^bold>P _ _" [1000, 65] 65) + and Eventually ("\<^bold>F _ _" [1000, 65] 65) + and Historically ("\<^bold>H _ _" [1000, 65] 65) + and Always ("\<^bold>G _ _" [1000, 65] 65) + and Since ("_ \<^bold>S _ _" [60,1000,60] 60) + and Until ("_ \<^bold>U _ _" [60,1000,60] 60) + +no_notation eval_trm ("_\_\" [70,89] 89) + and eval_trms ("_\<^bold>\_\<^bold>\" [70,89] 89) + and eval_trm_set ("_\_\" [70,89] 89) + and eval_trms_set ("_\<^bold>\_\<^bold>\" [70,89] 89) + and sat ("\_, _, _\ \ _" [56, 56, 56, 56] 55) + and Interval.interval ("\<^bold>[_,_\<^bold>]") + +end + +bundle MFOTL_notation begin + +notation Var ("\<^bold>v") + and Const ("\<^bold>c") + +notation TT ("\") + and FF ("\") + and Pred ("_ \ _" [85, 85] 85) + and Eq_Const ("_ \<^bold>\ _" [85, 85] 85) + and Neg ("\\<^sub>F _" [82] 82) + and And (infixr "\\<^sub>F" 80) + and Or (infixr "\\<^sub>F" 80) + and Imp (infixr "\\<^sub>F" 79) + and Iff (infixr "\\<^sub>F" 79) + and Exists ("\\<^sub>F_. _" [70,70] 70) + and Forall ("\\<^sub>F_. _" [70,70] 70) + and Prev ("\<^bold>Y _ _" [1000, 65] 65) + and Next ("\<^bold>X _ _" [1000, 65] 65) + and Once ("\<^bold>P _ _" [1000, 65] 65) + and Eventually ("\<^bold>F _ _" [1000, 65] 65) + and Historically ("\<^bold>H _ _" [1000, 65] 65) + and Always ("\<^bold>G _ _" [1000, 65] 65) + and Since ("_ \<^bold>S _ _" [60,1000,60] 60) + and Until ("_ \<^bold>U _ _" [60,1000,60] 60) + +notation eval_trm ("_\_\" [70,89] 89) + and eval_trms ("_\<^bold>\_\<^bold>\" [70,89] 89) + and eval_trm_set ("_\_\" [70,89] 89) + and eval_trms_set ("_\<^bold>\_\<^bold>\" [70,89] 89) + and sat ("\_, _, _\ \ _" [56, 56, 56, 56] 55) + and Interval.interval ("\<^bold>[_,_\<^bold>]") + +end + +unbundle MFOTL_no_notation + +(*<*) +end +(*>*) \ No newline at end of file diff --git a/thys/MFOTL_Checker/Monitor.thy b/thys/MFOTL_Checker/Monitor.thy new file mode 100644 --- /dev/null +++ b/thys/MFOTL_Checker/Monitor.thy @@ -0,0 +1,587 @@ +(*<*) +theory Monitor + imports Checker_Code +begin +(*>*) + +section \Unverified Explanation-Producing Monitoring Algorithm\ + +fun merge_part2_raw :: "('a \ 'b \ 'c) \ ('d set \ 'a) list \ ('d set \ 'b) list \ ('d set \ 'c) list" where + "merge_part2_raw f [] _ = []" +| "merge_part2_raw f ((P1, v1) # part1) part2 = + (let part12 = List.map_filter (\(P2, v2). if P1 \ P2 \ {} then Some(P1 \ P2, f v1 v2) else None) part2 in + let part2not1 = List.map_filter (\(P2, v2). if P2 - P1 \ {} then Some(P2 - P1, v2) else None) part2 in + part12 @ (merge_part2_raw f part1 part2not1))" + +fun merge_part3_raw :: "('a \ 'b \ 'c \ 'e) \ ('d set \ 'a) list \ ('d set \ 'b) list \ ('d set \ 'c) list \ ('d set \ 'e) list" where + "merge_part3_raw f [] _ _ = []" +| "merge_part3_raw f _ [] _ = []" +| "merge_part3_raw f _ _ [] = []" +| "merge_part3_raw f part1 part2 part3 = merge_part2_raw (\pt3 f'. f' pt3) part3 (merge_part2_raw f part1 part2)" + +lemma partition_on_empty_iff: + "partition_on X \

    \ \

    = {} \ X = {}" + "partition_on X \

    \ \

    \ {} \ X \ {}" + by (auto simp: partition_on_def) + +lemma wf_part_list_filter_inter: + defines "inP1 P1 f v1 part2 + \ List.map_filter (\(P2, v2). if P1 \ P2 \ {} then Some(P1 \ P2, f v1 v2) else None) part2" + assumes "partition_on X (set (map fst ((P1, v1) # part1)))" + and "partition_on X (set (map fst part2))" + shows "partition_on P1 (set (map fst (inP1 P1 f v1 part2)))" + and "distinct (map fst ((P1, v1) # part1)) \ distinct (map fst (part2)) \ + distinct (map fst (inP1 P1 f v1 part2))" +proof (rule partition_onI) + show "\ (set (map fst (inP1 P1 f v1 part2))) = P1" + proof - + have "\P2. (P1 \ P2 \ {} \ (\v2. (P2, v2) \ set part2) \ x \ P2) \ P1 \ P2 \ {}" + if "\ (fst ` set part2) = P1 \ \ (fst ` set part1)" and "x \ P1" for x + using that by (metis (no_types, lifting) Int_iff UN_iff Un_Int_eq(3) empty_iff prod.collapse) + with partition_onD1[OF assms(2)] partition_onD1[OF assms(3)] show ?thesis + by (auto simp: map_filter_def inP1_def split: if_splits) + qed + show "\A1 A2. A1 \ set (map fst (inP1 P1 f v1 part2)) \ + A2 \ set (map fst (inP1 P1 f v1 part2)) \ A1 \ A2 \ disjnt A1 A2" + using partition_onD2[OF assms(2)] partition_onD2[OF assms(3)] + by (force simp: disjnt_iff map_filter_def disjoint_def inP1_def split: if_splits) + show "{} \ set (map fst (inP1 P1 f v1 part2))" + using assms + by (auto simp: map_filter_def split: if_splits) + show "distinct (map fst ((P1, v1) # part1)) \ distinct (map fst part2) \ + distinct (map fst (inP1 P1 f v1 part2))" + using partition_onD2[OF assms(3), unfolded disjoint_def, simplified, rule_format] + by (clarsimp simp: inP1_def map_filter_def distinct_map inj_on_def Ball_def) blast +qed + +lemma wf_part_list_filter_minus: + defines "notinP2 P1 f v1 part2 + \ List.map_filter (\(P2, v2). if P2 - P1 \ {} then Some(P2 - P1, v2) else None) part2" + assumes "partition_on X (set (map fst ((P1, v1) # part1)))" + and "partition_on X (set (map fst part2))" + shows "partition_on (X - P1) (set (map fst (notinP2 P1 f v1 part2)))" + and "distinct (map fst ((P1, v1) # part1)) \ distinct (map fst (part2)) \ + distinct (map fst (notinP2 P1 f v1 part2))" +proof (rule partition_onI) + show "\ (set (map fst (notinP2 P1 f v1 part2))) = X - P1" + proof - + have "\P2. ((\x\P2. x \ P1) \ (\v2. (P2, v2) \ set part2)) \ (\x\P2. x \ P1) \ x \ P2" + if "\ (fst ` set part2) = P1 \ \ (fst ` set part1)" "x \ P1" "(P1', v1) \ set part1" "x \ P1'" for x P1' v1 + using that by (metis (no_types, lifting) UN_iff Un_iff fst_conv prod.collapse) + with partition_onD1[OF assms(2)] partition_onD1[OF assms(3)] show ?thesis + by (auto simp: map_filter_def subset_eq split_beta notinP2_def split: if_splits) + qed + show "\A1 A2. A1 \ set (map fst (notinP2 P1 f v1 part2)) \ + A2 \ set (map fst (notinP2 P1 f v1 part2)) \ A1 \ A2 \ disjnt A1 A2" + using partition_onD2[OF assms(3)] + by (auto simp: disjnt_def map_filter_def disjoint_def notinP2_def Ball_def Bex_def image_iff split: if_splits) + show "{} \ set (map fst (notinP2 P1 f v1 part2))" + using assms + by (auto simp: map_filter_def split: if_splits) + show "distinct (map fst ((P1, v1) # part1)) \ distinct (map fst part2) \ + distinct (map fst ((notinP2 P1 f v1 part2)))" + using partition_onD2[OF assms(3), unfolded disjoint_def] + by (clarsimp simp: notinP2_def map_filter_def distinct_map inj_on_def Ball_def Bex_def image_iff) blast +qed + +lemma wf_part_list_tail: + assumes "partition_on X (set (map fst ((P1, v1) # part1)))" + and "distinct (map fst ((P1, v1) # part1))" + shows "partition_on (X - P1) (set (map fst part1))" + and "distinct (map fst part1)" +proof (rule partition_onI) + show "\ (set (map fst part1)) = X - P1" + using partition_onD1[OF assms(1)] partition_onD2[OF assms(1)] assms(2) + by (auto simp: disjoint_def image_iff) + show "\A1 A2. A1 \ set (map fst part1) \ A2 \ set (map fst part1) \ A1 \ A2 \ disjnt A1 A2" + using partition_onD2[OF assms(1)] + by (clarsimp simp: disjnt_def disjoint_def) + (smt (verit, ccfv_SIG) Diff_disjoint Int_Diff Int_commute fst_conv) + show "{} \ set (map fst part1)" + using partition_onD3[OF assms(1)] + by (auto simp: map_filter_def split: if_splits) + show "distinct (map fst (part1))" + using assms(2) + by auto +qed + +lemma partition_on_append: "partition_on X (set xs) \ partition_on Y (set ys) \ X \ Y = {} \ + partition_on (X \ Y) (set (xs @ ys))" + by (auto simp: partition_on_def intro!: disjoint_union) + +lemma wf_part_list_merge_part2_raw: + "partition_on X (set (map fst part1)) \ distinct (map fst part1) \ + partition_on X (set (map fst part2)) \ distinct (map fst part2) \ + partition_on X (set (map fst (merge_part2_raw f part1 part2))) + \ distinct (map fst (merge_part2_raw f part1 part2))" +proof(induct f part1 part2 arbitrary: X rule: merge_part2_raw.induct) + case (2 f P1 v1 part1 part2) + let ?inP1 = "List.map_filter (\(P2, v2). if P1 \ P2 \ {} then Some (P1 \ P2, f v1 v2) else None) part2" + and ?notinP1 = "List.map_filter (\(P2, v2). if P2 - P1 \ {} then Some (P2 - P1, v2) else None) part2" + have "P1 \ X = X" + using "2.prems" + by (auto simp: partition_on_def) + have wf_part1: "partition_on (X - P1) (set (map fst part1))" + "distinct (map fst part1)" + using wf_part_list_tail "2.prems" by auto + moreover have wf_notinP1: "partition_on (X - P1) (set (map fst ?notinP1))" + "distinct (map fst (?notinP1))" + using wf_part_list_filter_minus[OF 2(2)[THEN conjunct1]] + "2.prems" by auto + ultimately have IH: "partition_on (X - P1) (set (map fst (merge_part2_raw f part1 (?notinP1))))" + "distinct (map fst (merge_part2_raw f part1 (?notinP1)))" + using "2.hyps"[OF refl refl] by auto + moreover have wf_inP1: "partition_on P1 (set (map fst ?inP1))" "distinct (map fst ?inP1)" + using wf_part_list_filter_inter[OF 2(2)[THEN conjunct1]] + "2.prems" by auto + moreover have "(fst ` set ?inP1) \ (fst ` set (merge_part2_raw f part1 (?notinP1))) = {}" + using IH(1)[THEN partition_onD1] + by (fastforce simp: map_filter_def split: prod.splits if_splits) + ultimately show ?case + using partition_on_append[OF wf_inP1(1) IH(1)] \P1 \ X = X\ wf_inP1(2) IH(2) + by simp +qed simp + +lemma wf_part_list_merge_part3_raw: + "partition_on X (set (map fst part1)) \ distinct (map fst part1) \ + partition_on X (set (map fst part2)) \ distinct (map fst part2) \ + partition_on X (set (map fst part3)) \ distinct (map fst part3) \ + partition_on X (set (map fst (merge_part3_raw f part1 part2 part3))) + \ distinct (map fst (merge_part3_raw f part1 part2 part3))" +proof(induct f part1 part2 part3 arbitrary: X rule: merge_part3_raw.induct) + case (4 f v va vb vc vd ve) + have "partition_on X (set (map fst (v # va))) \ distinct (map fst (vb # vc))" + using 4 by blast + moreover have "partition_on X (set (map fst (vb # vc))) \ distinct (map fst (vb # vc))" + using 4 by blast + ultimately have "partition_on X (set (map fst (merge_part2_raw f (v # va) (vb # vc)))) + \ distinct (map fst (merge_part2_raw f (v # va) (vb # vc)))" + using wf_part_list_merge_part2_raw[of X "(v # va)" "(vb # vc)" f] 4 + by fastforce + moreover have "partition_on X (set (map fst (vd # ve))) \ distinct (map fst (vd # ve))" + using 4 by blast + ultimately show ?case + using wf_part_list_merge_part2_raw[of X "(vd # ve)" "(merge_part2_raw f (v # va) (vb # vc))" "(\pt3 f'. f' pt3)"] + by simp +qed auto + +lift_definition merge_part2 :: "('a \ 'a \ 'a) \ ('d, 'a) part \ ('d, 'a) part \ ('d, 'a) part" is merge_part2_raw + by (rule wf_part_list_merge_part2_raw) + +lift_definition merge_part3 :: "('a \ 'a \ 'a \ 'a) \ ('d, 'a) part \ ('d, 'a) part \ ('d, 'a) part \ ('d, 'a) part" is merge_part3_raw + by (rule wf_part_list_merge_part3_raw) + +definition proof_app :: "('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof" (infixl "\" 65) where + "p \ q = (case (p, q) of + (Inl (SHistorically i li sps), Inl q) \ Inl (SHistorically (i+1) li (sps @ [q])) + | (Inl (SAlways i hi sps), Inl q) \ Inl (SAlways (i-1) hi (q # sps)) + | (Inl (SSince sp2 sp1s), Inl q) \ Inl (SSince sp2 (sp1s @ [q])) + | (Inl (SUntil sp1s sp2), Inl q) \ Inl (SUntil (q # sp1s) sp2) + | (Inr (VSince i vp1 vp2s), Inr q) \ Inr (VSince (i+1) vp1 (vp2s @ [q])) + | (Inr (VOnce i li vps), Inr q) \ Inr (VOnce (i+1) li (vps @ [q])) + | (Inr (VEventually i hi vps), Inr q) \ Inr (VEventually (i-1) hi (q # vps)) + | (Inr (VSinceInf i li vp2s), Inr q) \ Inr (VSinceInf (i+1) li (vp2s @ [q])) + | (Inr (VUntil i vp2s vp1), Inr q) \ Inr (VUntil (i-1) (q # vp2s) vp1) + | (Inr (VUntilInf i hi vp2s), Inr q) \ Inr (VUntilInf (i-1) hi (q # vp2s)))" + +definition proof_incr :: "('n, 'd) proof \ ('n, 'd) proof" where + "proof_incr p = (case p of + Inl (SOnce i sp) \ Inl (SOnce (i+1) sp) + | Inl (SEventually i sp) \ Inl (SEventually (i-1) sp) + | Inl (SHistorically i li sps) \ Inl (SHistorically (i+1) li sps) + | Inl (SAlways i hi sps) \ Inl (SAlways (i-1) hi sps) + | Inr (VSince i vp1 vp2s) \ Inr (VSince (i+1) vp1 vp2s) + | Inr (VOnce i li vps) \ Inr (VOnce (i+1) li vps) + | Inr (VEventually i hi vps) \ Inr (VEventually (i-1) hi vps) + | Inr (VHistorically i vp) \ Inr (VHistorically (i+1) vp) + | Inr (VAlways i vp) \ Inr (VAlways (i-1) vp) + | Inr (VSinceInf i li vp2s) \ Inr (VSinceInf (i+1) li vp2s) + | Inr (VUntil i vp2s vp1) \ Inr (VUntil (i-1) vp2s vp1) + | Inr (VUntilInf i hi vp2s) \ Inr (VUntilInf (i-1) hi vp2s))" + +definition min_list_wrt :: "(('n, 'd) proof \ ('n, 'd) proof \ bool) \ ('n, 'd) proof list \ ('n, 'd) proof" where + "min_list_wrt r xs = hd [x \ xs. \y \ set xs. r x y]" + +definition do_neg :: "('n, 'd) proof \ ('n, 'd) proof list" where + "do_neg p = (case p of + Inl sp \ [Inr (VNeg sp)] +| Inr vp \ [Inl (SNeg vp)])" + +definition do_or :: "('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_or p1 p2 = (case (p1, p2) of + (Inl sp1, Inl sp2) \ [Inl (SOrL sp1), Inl (SOrR sp2)] +| (Inl sp1, Inr _ ) \ [Inl (SOrL sp1)] +| (Inr _ , Inl sp2) \ [Inl (SOrR sp2)] +| (Inr vp1, Inr vp2) \ [Inr (VOr vp1 vp2)])" + +definition do_and :: "('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_and p1 p2 = (case (p1, p2) of + (Inl sp1, Inl sp2) \ [Inl (SAnd sp1 sp2)] +| (Inl _ , Inr vp2) \ [Inr (VAndR vp2)] +| (Inr vp1, Inl _ ) \ [Inr (VAndL vp1)] +| (Inr vp1, Inr vp2) \ [Inr (VAndL vp1), Inr (VAndR vp2)])" + +definition do_imp :: "('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_imp p1 p2 = (case (p1, p2) of + (Inl _ , Inl sp2) \ [Inl (SImpR sp2)] +| (Inl sp1, Inr vp2) \ [Inr (VImp sp1 vp2)] +| (Inr vp1, Inl sp2) \ [Inl (SImpL vp1), Inl (SImpR sp2)] +| (Inr vp1, Inr _ ) \ [Inl (SImpL vp1)])" + +definition do_iff :: "('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_iff p1 p2 = (case (p1, p2) of + (Inl sp1, Inl sp2) \ [Inl (SIffSS sp1 sp2)] +| (Inl sp1, Inr vp2) \ [Inr (VIffSV sp1 vp2)] +| (Inr vp1, Inl sp2) \ [Inr (VIffVS vp1 sp2)] +| (Inr vp1, Inr vp2) \ [Inl (SIffVV vp1 vp2)])" + +definition do_exists :: "'n \ ('n, 'd::{default,linorder}) proof + ('d, ('n, 'd) proof) part \ ('n, 'd) proof list" where + "do_exists x p_part = (case p_part of + Inl p \ (case p of + Inl sp \ [Inl (SExists x default sp)] + | Inr vp \ [Inr (VExists x (trivial_part vp))]) +| Inr part \ (if (\x\Vals part. isl x) then + map (\(D,p). map_sum (SExists x (Min D)) id p) (filter (\(_, p). isl p) (subsvals part)) + else + [Inr (VExists x (map_part projr part))]))" + +definition do_forall :: "'n \ ('n, 'd::{default,linorder}) proof + ('d, ('n, 'd) proof) part \ ('n, 'd) proof list" where + "do_forall x p_part = (case p_part of + Inl p \ (case p of + Inl sp \ [Inl (SForall x (trivial_part sp))] + | Inr vp \ [Inr (VForall x default vp)]) +| Inr part \ (if (\x\Vals part. isl x) then + [Inl (SForall x (map_part projl part))] + else + map (\(D,p). map_sum id (VForall x (Min D)) p) (filter (\(_, p). \isl p) (subsvals part))))" + +definition do_prev :: "nat \ \ \ nat \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_prev i I t p = (case (p, t < left I) of + (Inl _ , True) \ [Inr (VPrevOutL i)] +| (Inl sp, False) \ (if mem t I then [Inl (SPrev sp)] else [Inr (VPrevOutR i)]) +| (Inr vp, True) \ [Inr (VPrev vp), Inr (VPrevOutL i)] +| (Inr vp, False) \ (if mem t I then [Inr (VPrev vp)] else [Inr (VPrev vp), Inr (VPrevOutR i)]))" + +definition do_next :: "nat \ \ \ nat \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_next i I t p = (case (p, t < left I) of + (Inl _ , True) \ [Inr (VNextOutL i)] +| (Inl sp, False) \ (if mem t I then [Inl (SNext sp)] else [Inr (VNextOutR i)]) +| (Inr vp, True) \ [Inr (VNext vp), Inr (VNextOutL i)] +| (Inr vp, False) \ (if mem t I then [Inr (VNext vp)] else [Inr (VNext vp), Inr (VNextOutR i)]))" + +definition do_once_base :: "nat \ nat \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_once_base i a p' = (case (p', a = 0) of + (Inl sp', True) \ [Inl (SOnce i sp')] +| (Inr vp', True) \ [Inr (VOnce i i [vp'])] +| ( _ , False) \ [Inr (VOnce i i [])])" + +definition do_once :: "nat \ nat \ ('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_once i a p p' = (case (p, a = 0, p') of + (Inl sp, True, Inr _ ) \ [Inl (SOnce i sp)] +| (Inl sp, True, Inl (SOnce _ sp')) \ [Inl (SOnce i sp'), Inl (SOnce i sp)] +| (Inl _ , False, Inl (SOnce _ sp')) \ [Inl (SOnce i sp')] +| (Inl _ , False, Inr (VOnce _ li vps')) \ [Inr (VOnce i li vps')] +| (Inr _ , True, Inl (SOnce _ sp')) \ [Inl (SOnce i sp')] +| (Inr vp, True, Inr vp') \ [(Inr vp') \ (Inr vp)] +| (Inr _ , False, Inl (SOnce _ sp')) \ [Inl (SOnce i sp')] +| (Inr _ , False, Inr (VOnce _ li vps')) \ [Inr (VOnce i li vps')])" + +definition do_eventually_base :: "nat \ nat \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_eventually_base i a p' = (case (p', a = 0) of + (Inl sp', True) \ [Inl (SEventually i sp')] +| (Inr vp', True) \ [Inr (VEventually i i [vp'])] +| ( _ , False) \ [Inr (VEventually i i [])])" + +definition do_eventually :: "nat \ nat \ ('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_eventually i a p p' = (case (p, a = 0, p') of + (Inl sp, True, Inr _ ) \ [Inl (SEventually i sp)] +| (Inl sp, True, Inl (SEventually _ sp')) \ [Inl (SEventually i sp'), Inl (SEventually i sp)] +| (Inl _ , False, Inl (SEventually _ sp')) \ [Inl (SEventually i sp')] +| (Inl _ , False, Inr (VEventually _ hi vps')) \ [Inr (VEventually i hi vps')] +| (Inr _ , True, Inl (SEventually _ sp')) \ [Inl (SEventually i sp')] +| (Inr vp, True, Inr vp') \ [(Inr vp') \ (Inr vp)] +| (Inr _ , False, Inl (SEventually _ sp')) \ [Inl (SEventually i sp')] +| (Inr _ , False, Inr (VEventually _ hi vps')) \ [Inr (VEventually i hi vps')])" + +definition do_historically_base :: "nat \ nat \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_historically_base i a p' = (case (p', a = 0) of + (Inl sp', True) \ [Inl (SHistorically i i [sp'])] +| (Inr vp', True) \ [Inr (VHistorically i vp')] +| ( _ , False) \ [Inl (SHistorically i i [])])" + +definition do_historically :: "nat \ nat \ ('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_historically i a p p' = (case (p, a = 0, p') of + (Inl _ , True, Inr (VHistorically _ vp')) \ [Inr (VHistorically i vp')] +| (Inl sp, True, Inl sp') \ [(Inl sp') \ (Inl sp)] +| (Inl _ , False, Inl (SHistorically _ li sps')) \ [Inl (SHistorically i li sps')] +| (Inl _ , False, Inr (VHistorically _ vp')) \ [Inr (VHistorically i vp')] +| (Inr vp, True, Inl _ ) \ [Inr (VHistorically i vp)] +| (Inr vp, True, Inr (VHistorically _ vp')) \ [Inr (VHistorically i vp), Inr (VHistorically i vp')] +| (Inr _ , False, Inl (SHistorically _ li sps')) \ [Inl (SHistorically i li sps')] +| (Inr _ , False, Inr (VHistorically _ vp')) \ [Inr (VHistorically i vp')])" + +definition do_always_base :: "nat \ nat \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_always_base i a p' = (case (p', a = 0) of + (Inl sp', True) \ [Inl (SAlways i i [sp'])] +| (Inr vp', True) \ [Inr (VAlways i vp')] +| ( _ , False) \ [Inl (SAlways i i [])])" + +definition do_always :: "nat \ nat \ ('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_always i a p p' = (case (p, a = 0, p') of + (Inl _ , True, Inr (VAlways _ vp')) \ [Inr (VAlways i vp')] +| (Inl sp, True, Inl sp') \ [(Inl sp') \ (Inl sp)] +| (Inl _ , False, Inl (SAlways _ hi sps')) \ [Inl (SAlways i hi sps')] +| (Inl _ , False, Inr (VAlways _ vp')) \ [Inr (VAlways i vp')] +| (Inr vp, True, Inl _ ) \ [Inr (VAlways i vp)] +| (Inr vp, True, Inr (VAlways _ vp')) \ [Inr (VAlways i vp), Inr (VAlways i vp')] +| (Inr _ , False, Inl (SAlways _ hi sps')) \ [Inl (SAlways i hi sps')] +| (Inr _ , False, Inr (VAlways _ vp')) \ [Inr (VAlways i vp')])" + +definition do_since_base :: "nat \ nat \ ('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_since_base i a p1 p2 = (case (p1, p2, a = 0) of + ( _ , Inl sp2, True) \ [Inl (SSince sp2 [])] +| (Inl _ , _ , False) \ [Inr (VSinceInf i i [])] +| (Inl _ , Inr vp2, True) \ [Inr (VSinceInf i i [vp2])] +| (Inr vp1, _ , False) \ [Inr (VSince i vp1 []), Inr (VSinceInf i i [])] +| (Inr vp1, Inr sp2, True) \ [Inr (VSince i vp1 [sp2]), Inr (VSinceInf i i [sp2])])" + +definition do_since :: "nat \ nat \ ('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_since i a p1 p2 p' = (case (p1, p2, a = 0, p') of + (Inl sp1, Inr _ , True, Inl sp') \ [(Inl sp') \ (Inl sp1)] +| (Inl sp1, _ , False, Inl sp') \ [(Inl sp') \ (Inl sp1)] +| (Inl sp1, Inl sp2, True, Inl sp') \ [(Inl sp') \ (Inl sp1), Inl (SSince sp2 [])] +| (Inl _ , Inr vp2, True, Inr (VSinceInf _ _ _ )) \ [p' \ (Inr vp2)] +| (Inl _ , _ , False, Inr (VSinceInf _ li vp2s')) \ [Inr (VSinceInf i li vp2s')] +| (Inl _ , Inr vp2, True, Inr (VSince _ _ _ )) \ [p' \ (Inr vp2)] +| (Inl _ , _ , False, Inr (VSince _ vp1' vp2s')) \ [Inr (VSince i vp1' vp2s')] +| (Inr vp1, Inr vp2, True, Inl _ ) \ [Inr (VSince i vp1 [vp2])] +| (Inr vp1, _ , False, Inl _ ) \ [Inr (VSince i vp1 [])] +| (Inr _ , Inl sp2, True, Inl _ ) \ [Inl (SSince sp2 [])] +| (Inr vp1, Inr vp2, True, Inr (VSinceInf _ _ _ )) \ [Inr (VSince i vp1 [vp2]), p' \ (Inr vp2)] +| (Inr vp1, _, False, Inr (VSinceInf _ li vp2s')) \ [Inr (VSince i vp1 []), Inr (VSinceInf i li vp2s')] +| ( _ , Inl sp2, True, Inr (VSinceInf _ _ _ )) \ [Inl (SSince sp2 [])] +| (Inr vp1, Inr vp2, True, Inr (VSince _ _ _ )) \ [Inr (VSince i vp1 [vp2]), p' \ (Inr vp2)] +| (Inr vp1, _ , False, Inr (VSince _ vp1' vp2s')) \ [Inr (VSince i vp1 []), Inr (VSince i vp1' vp2s')] +| ( _ , Inl vp2, True, Inr (VSince _ _ _ )) \ [Inl (SSince vp2 [])])" + +definition do_until_base :: "nat \ nat \ ('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_until_base i a p1 p2 = (case (p1, p2, a = 0) of + ( _ , Inl sp2, True) \ [Inl (SUntil [] sp2)] +| (Inl sp1, _ , False) \ [Inr (VUntilInf i i [])] +| (Inl sp1, Inr vp2, True) \ [Inr (VUntilInf i i [vp2])] +| (Inr vp1, _ , False) \ [Inr (VUntil i [] vp1), Inr (VUntilInf i i [])] +| (Inr vp1, Inr vp2, True) \ [Inr (VUntil i [vp2] vp1), Inr (VUntilInf i i [vp2])])" + +definition do_until :: "nat \ nat \ ('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof list" where + "do_until i a p1 p2 p' = (case (p1, p2, a = 0, p') of + (Inl sp1, Inr _ , True, Inl (SUntil _ _ )) \ [p' \ (Inl sp1)] +| (Inl sp1, _ , False, Inl (SUntil _ _ )) \ [p' \ (Inl sp1)] +| (Inl sp1, Inl sp2, True, Inl (SUntil _ _ )) \ [p' \ (Inl sp1), Inl (SUntil [] sp2)] +| (Inl _ , Inr vp2, True, Inr (VUntilInf _ _ _ )) \ [p' \ (Inr vp2)] +| (Inl _ , _ , False, Inr (VUntilInf _ hi vp2s')) \ [Inr (VUntilInf i hi vp2s')] +| (Inl _ , Inr vp2, True, Inr (VUntil _ _ _ )) \ [p' \ (Inr vp2)] +| (Inl _ , _ , False, Inr (VUntil _ vp2s' vp1')) \ [Inr (VUntil i vp2s' vp1')] +| (Inr vp1, Inr vp2, True, Inl (SUntil _ _ )) \ [Inr (VUntil i [vp2] vp1)] +| (Inr vp1, _ , False, Inl (SUntil _ _ )) \ [Inr (VUntil i [] vp1)] +| (Inr vp1, Inl sp2, True, Inl (SUntil _ _ )) \ [Inl (SUntil [] sp2)] +| (Inr vp1, Inr vp2, True, Inr (VUntilInf _ _ _ )) \ [Inr (VUntil i [vp2] vp1), p' \ (Inr vp2)] +| (Inr vp1, _ , False, Inr (VUntilInf _ hi vp2s')) \ [Inr (VUntil i [] vp1), Inr (VUntilInf i hi vp2s')] +| ( _ , Inl sp2, True, Inr (VUntilInf _ hi vp2s')) \ [Inl (SUntil [] sp2)] +| (Inr vp1, Inr vp2, True, Inr (VUntil _ _ _ )) \ [Inr (VUntil i [vp2] vp1), p' \ (Inr vp2)] +| (Inr vp1, _ , False, Inr (VUntil _ vp2s' vp1')) \ [Inr (VUntil i [] vp1), Inr (VUntil i vp2s' vp1')] +| ( _ , Inl sp2, True, Inr (VUntil _ _ _ )) \ [Inl (SUntil [] sp2)])" + +fun match :: "('n, 'd) Formula.trm list \ 'd list \ ('n \ 'd) option" where + "match [] [] = Some Map.empty" +| "match (Formula.Const x # ts) (y # ys) = (if x = y then match ts ys else None)" +| "match (Formula.Var x # ts) (y # ys) = (case match ts ys of + None \ None + | Some f \ (case f x of + None \ Some (f(x \ y)) + | Some z \ if y = z then Some f else None))" +| "match _ _ = None" + +fun pdt_of :: "nat \ 'n \ ('n, 'd :: linorder) Formula.trm list \ 'n list \ ('n \ 'd) list \ ('n, 'd) expl" where + "pdt_of i r ts [] V = (if List.null V then Leaf (Inr (VPred i r ts)) else Leaf (Inl (SPred i r ts)))" +| "pdt_of i r ts (x # vs) V = + (let ds = remdups (List.map_filter (\v. v x) V); + part = tabulate ds (\d. pdt_of i r ts vs (filter (\v. v x = Some d) V)) (pdt_of i r ts vs []) + in Node x part)" + +fun "apply_pdt1" :: "'n list \ (('n, 'd) proof \ ('n, 'd) proof) \ ('n, 'd) expl \ ('n, 'd) expl" where + "apply_pdt1 vs f (Leaf pt) = Leaf (f pt)" +| "apply_pdt1 (z # vs) f (Node x part) = + (if x = z then + Node x (map_part (\expl. apply_pdt1 vs f expl) part) + else + apply_pdt1 vs f (Node x part))" +| "apply_pdt1 [] _ (Node _ _) = undefined" + +fun "apply_pdt2" :: "'n list \ (('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof) \ ('n, 'd) expl \ ('n, 'd) expl \ ('n, 'd) expl" where + "apply_pdt2 vs f (Leaf pt1) (Leaf pt2) = Leaf (f pt1 pt2)" +| "apply_pdt2 vs f (Leaf pt1) (Node x part2) = Node x (map_part (apply_pdt1 vs (f pt1)) part2)" +| "apply_pdt2 vs f (Node x part1) (Leaf pt2) = Node x (map_part (apply_pdt1 vs (\pt1. f pt1 pt2)) part1)" +| "apply_pdt2 (z # vs) f (Node x part1) (Node y part2) = + (if x = z \ y = z then + Node z (merge_part2 (apply_pdt2 vs f) part1 part2) + else if x = z then + Node x (map_part (\expl1. apply_pdt2 vs f expl1 (Node y part2)) part1) + else if y = z then + Node y (map_part (\expl2. apply_pdt2 vs f (Node x part1) expl2) part2) + else + apply_pdt2 vs f (Node x part1) (Node y part2))" +| "apply_pdt2 [] _ (Node _ _) (Node _ _) = undefined" + +fun "apply_pdt3" :: "'n list \ (('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof \ ('n, 'd) proof) \ ('n, 'd) expl \ ('n, 'd) expl \ ('n, 'd) expl \ ('n, 'd) expl" where + "apply_pdt3 vs f (Leaf pt1) (Leaf pt2) (Leaf pt3) = Leaf (f pt1 pt2 pt3)" +| "apply_pdt3 vs f (Leaf pt1) (Leaf pt2) (Node x part3) = Node x (map_part (apply_pdt2 vs (f pt1) (Leaf pt2)) part3)" +| "apply_pdt3 vs f (Leaf pt1) (Node x part2) (Leaf pt3) = Node x (map_part (apply_pdt2 vs (\pt2. f pt1 pt2) (Leaf pt3)) part2)" +| "apply_pdt3 vs f (Node x part1) (Leaf pt2) (Leaf pt3) = Node x (map_part (apply_pdt2 vs (\pt1. f pt1 pt2) (Leaf pt3)) part1)" +| "apply_pdt3 (w # vs) f (Leaf pt1) (Node y part2) (Node z part3) = + (if y = w \ z = w then + Node w (merge_part2 (apply_pdt2 vs (f pt1)) part2 part3) + else if y = w then + Node y (map_part (\expl2. apply_pdt2 vs (f pt1) expl2 (Node z part3)) part2) + else if z = w then + Node z (map_part (\expl3. apply_pdt2 vs (f pt1) (Node y part2) expl3) part3) + else + apply_pdt3 vs f (Leaf pt1) (Node y part2) (Node z part3))" +| "apply_pdt3 (w # vs) f (Node x part1) (Node y part2) (Leaf pt3) = + (if x = w \ y = w then + Node w (merge_part2 (apply_pdt2 vs (\pt1 pt2. f pt1 pt2 pt3)) part1 part2) + else if x = w then + Node x (map_part (\expl1. apply_pdt2 vs (\pt1 pt2. f pt1 pt2 pt3) expl1 (Node y part2)) part1) + else if y = w then + Node y (map_part (\expl2. apply_pdt2 vs (\pt1 pt2. f pt1 pt2 pt3) (Node x part1) expl2) part2) + else + apply_pdt3 vs f (Node x part1) (Node y part2) (Leaf pt3))" +| "apply_pdt3 (w # vs) f (Node x part1) (Leaf pt2) (Node z part3) = + (if x = w \ z = w then + Node w (merge_part2 (apply_pdt2 vs (\pt1. f pt1 pt2)) part1 part3) + else if x = w then + Node x (map_part (\expl1. apply_pdt2 vs (\pt1. f pt1 pt2) expl1 (Node z part3)) part1) + else if z = w then + Node z (map_part (\expl3. apply_pdt2 vs (\pt1. f pt1 pt2) (Node x part1) expl3) part3) + else + apply_pdt3 vs f (Node x part1) (Leaf pt2) (Node z part3))" +| "apply_pdt3 (w # vs) f (Node x part1) (Node y part2) (Node z part3) = + (if x = w \ y = w \ z = w then + Node z (merge_part3 (apply_pdt3 vs f) part1 part2 part3) + else if x = w \ y = w then + Node w (merge_part2 (apply_pdt3 vs (\pt3 pt1 pt2. f pt1 pt2 pt3) (Node z part3)) part1 part2) + else if x = w \ z = w then + Node w (merge_part2 (apply_pdt3 vs (\pt2 pt1 pt3. f pt1 pt2 pt3) (Node y part2)) part1 part3) + else if y = w \ z = w then + Node w (merge_part2 (apply_pdt3 vs (\pt1. f pt1) (Node x part1)) part2 part3) + else if x = w then + Node x (map_part (\expl1. apply_pdt3 vs f expl1 (Node y part2) (Node z part3)) part1) + else if y = w then + Node y (map_part (\expl2. apply_pdt3 vs f (Node x part1) expl2 (Node z part3)) part2) + else if z = w then + Node z (map_part (\expl3. apply_pdt3 vs f (Node x part1) (Node y part2) expl3) part3) + else + apply_pdt3 vs f (Node x part1) (Node y part2) (Node z part3))" +| "apply_pdt3 [] _ _ _ _ = undefined" + +fun "hide_pdt" :: "'n list \ (('n, 'd) proof + ('d, ('n, 'd) proof) part \ ('n, 'd) proof) \ ('n, 'd) expl \ ('n, 'd) expl" where + "hide_pdt vs f (Leaf pt) = Leaf (f (Inl pt))" +| "hide_pdt [x] f (Node y part) = Leaf (f (Inr (map_part unleaf part)))" +| "hide_pdt (x # xs) f (Node y part) = + (if x = y then + Node y (map_part (hide_pdt xs f) part) + else + hide_pdt xs f (Node y part))" +| "hide_pdt [] _ _ = undefined" + +context + fixes \ :: "('n, 'd :: {default, linorder}) trace" and + cmp :: "('n, 'd) proof \ ('n, 'd) proof \ bool" +begin + +function (sequential) eval :: "'n list \ nat \ ('n, 'd) Formula.formula \ ('n, 'd) expl" where + "eval vs i Formula.TT = Leaf (Inl (STT i))" +| "eval vs i Formula.FF = Leaf (Inr (VFF i))" +| "eval vs i (Eq_Const x c) = Node x (tabulate [c] (\c. Leaf (Inl (SEq_Const i x c))) (Leaf (Inr (VEq_Const i x c))))" +| "eval vs i (Formula.Pred r ts) = + (pdt_of i r ts (filter (\x. x \ Formula.fv (Formula.Pred r ts)) vs) (List.map_filter (match ts) (sorted_list_of_set (snd ` {rd \ \ \ i. fst rd = r}))))" +| "eval vs i (Formula.Neg \) = apply_pdt1 vs (\p. min_list_wrt cmp (do_neg p)) (eval vs i \)" +| "eval vs i (Formula.Or \ \) = apply_pdt2 vs (\p1 p2. min_list_wrt cmp (do_or p1 p2)) (eval vs i \) (eval vs i \)" +| "eval vs i (Formula.And \ \) = apply_pdt2 vs (\p1 p2. min_list_wrt cmp (do_and p1 p2)) (eval vs i \) (eval vs i \)" +| "eval vs i (Formula.Imp \ \) = apply_pdt2 vs (\p1 p2. min_list_wrt cmp (do_imp p1 p2)) (eval vs i \) (eval vs i \)" +| "eval vs i (Formula.Iff \ \) = apply_pdt2 vs (\p1 p2. min_list_wrt cmp (do_iff p1 p2)) (eval vs i \) (eval vs i \)" +| "eval vs i (Formula.Exists x \) = hide_pdt (vs @ [x]) (\p. min_list_wrt cmp (do_exists x p)) (eval (vs @ [x]) i \)" +| "eval vs i (Formula.Forall x \) = hide_pdt (vs @ [x]) (\p. min_list_wrt cmp (do_forall x p)) (eval (vs @ [x]) i \)" +| "eval vs i (Formula.Prev I \) = (if i = 0 then Leaf (Inr VPrevZ) + else apply_pdt1 vs (\p. min_list_wrt cmp (do_prev i I (\ \ i) p)) (eval vs (i-1) \))" +| "eval vs i (Formula.Next I \) = apply_pdt1 vs (\l. min_list_wrt cmp (do_next i I (\ \ (i+1)) l)) (eval vs (i+1) \)" +| "eval vs i (Formula.Once I \) = + (if \ \ i < \ \ 0 + left I then Leaf (Inr (VOnceOut i)) + else (let expl = eval vs i \ in + (if i = 0 then + apply_pdt1 vs (\p. min_list_wrt cmp (do_once_base 0 0 p)) expl + else (if right I \ enat (\ \ i) then + apply_pdt2 vs (\p p'. min_list_wrt cmp (do_once i (left I) p p')) expl + (eval vs (i-1) (Formula.Once (subtract (\ \ i) I) \)) + else apply_pdt1 vs (\p. min_list_wrt cmp (do_once_base i (left I) p)) expl))))" +| "eval vs i (Formula.Historically I \) = + (if \ \ i < \ \ 0 + left I then Leaf (Inl (SHistoricallyOut i)) + else (let expl = eval vs i \ in + (if i = 0 then + apply_pdt1 vs (\p. min_list_wrt cmp (do_historically_base 0 0 p)) expl + else (if right I \ enat (\ \ i) then + apply_pdt2 vs (\p p'. min_list_wrt cmp (do_historically i (left I) p p')) expl + (eval vs (i-1) (Formula.Historically (subtract (\ \ i) I) \)) + else apply_pdt1 vs (\p. min_list_wrt cmp (do_historically_base i (left I) p)) expl))))" +| "eval vs i (Formula.Eventually I \) = + (let expl = eval vs i \ in + (if right I = \ then undefined + else (if right I \ enat (\ \ (i+1)) then + apply_pdt2 vs (\p p'. min_list_wrt cmp (do_eventually i (left I) p p')) expl + (eval vs (i+1) (Formula.Eventually (subtract (\ \ (i+1)) I) \)) + else apply_pdt1 vs (\p. min_list_wrt cmp (do_eventually_base i (left I) p)) expl)))" +| "eval vs i (Formula.Always I \) = + (let expl = eval vs i \ in + (if right I = \ then undefined + else (if right I \ enat (\ \ (i+1)) then + apply_pdt2 vs (\p p'. min_list_wrt cmp (do_always i (left I) p p')) expl + (eval vs (i+1) (Formula.Always (subtract (\ \ (i+1)) I) \)) + else apply_pdt1 vs (\p. min_list_wrt cmp (do_always_base i (left I) p)) expl)))" +| "eval vs i (Formula.Since \ I \) = + (if \ \ i < \ \ 0 + left I then Leaf (Inr (VSinceOut i)) + else (let expl1 = eval vs i \ in + let expl2 = eval vs i \ in + (if i = 0 then + apply_pdt2 vs (\p1 p2. min_list_wrt cmp (do_since_base 0 0 p1 p2)) expl1 expl2 + else (if right I \ enat (\ \ i) then + apply_pdt3 vs (\p1 p2 p'. min_list_wrt cmp (do_since i (left I) p1 p2 p')) expl1 expl2 + (eval vs (i-1) (Formula.Since \ (subtract (\ \ i) I) \)) + else apply_pdt2 vs (\p1 p2. min_list_wrt cmp (do_since_base i (left I) p1 p2)) expl1 expl2))))" +| "eval vs i (Formula.Until \ I \) = + (let expl1 = eval vs i \ in + let expl2 = eval vs i \ in + (if right I = \ then undefined + else (if right I \ enat (\ \ (i+1)) then + apply_pdt3 vs (\p1 p2 p'. min_list_wrt cmp (do_until i (left I) p1 p2 p')) expl1 expl2 + (eval vs (i+1) (Formula.Until \ (subtract (\ \ (i+1)) I) \)) + else apply_pdt2 vs (\p1 p2. min_list_wrt cmp (do_until_base i (left I) p1 p2)) expl1 expl2)))" + by pat_completeness auto + +fun dist where + "dist i (Formula.Once _ _) = i" +| "dist i (Formula.Historically _ _) = i" +| "dist i (Formula.Eventually I _) = LTP \ (case right I of \ \ 0 | enat b \ (\ \ i + b)) - i" +| "dist i (Formula.Always I _) = LTP \ (case right I of \ \ 0 | enat b \ (\ \ i + b)) - i" +| "dist i (Formula.Since _ _ _) = i" +| "dist i (Formula.Until _ I _) = LTP \ (case right I of \ \ 0 | enat b \ (\ \ i + b)) - i" +| "dist _ _ = undefined" + +lemma i_less_LTP: "\ \ (Suc i) \ b + \ \ i \ i < LTP \ (b + \ \ i)" + by (metis Suc_le_lessD i_le_LTPi_add le_iff_add) + +termination eval + by (relation "measures [\(_, _, \). size \, \(_, i, \). dist i \]") + (auto simp: add.commute le_diff_conv i_less_LTP intro!: diff_less_mono2) + +end + +end \ No newline at end of file diff --git a/thys/MFOTL_Checker/Partition.thy b/thys/MFOTL_Checker/Partition.thy new file mode 100644 --- /dev/null +++ b/thys/MFOTL_Checker/Partition.thy @@ -0,0 +1,297 @@ +(*<*) +theory Partition + imports "HOL-Library.Disjoint_Sets" +begin +(*>*) + +section \Valued Partitions\ + +lemma part_list_set_eq_aux1: + assumes + "\ij j \ fst (xs ! i) \ fst (xs ! j) = {}" + "{} \ fst ` set xs" + shows "disjoint (fst ` set xs) \ distinct (map fst xs)" +proof - + from assms(1) have "disjoint (fst ` set xs)" + by (metis disjnt_def in_set_conv_nth pairwise_imageI) + moreover have "distinct (map fst xs)" + using assms + by (smt (verit) distinct_conv_nth image_iff inf.idem + length_map nth_map nth_mem) + ultimately show ?thesis + by blast +qed + +lemma part_list_set_eq_aux2: + assumes + "disjoint (fst ` set xs)" + "distinct (map fst xs)" + "i < length xs" + "j < length xs" + "i \ j" + shows "fst (xs ! i) \ fst (xs ! j) = {}" +proof - + from assms have "fst (xs ! i) \ fst ` set xs" + and "fst (xs ! j) \ fst ` set xs" + by auto + moreover have "fst (xs ! i) \ fst (xs ! j)" + using assms(2-) nth_eq_iff_index_eq + by fastforce + ultimately show ?thesis + using assms(1) unfolding disjoint_def + by blast +qed + +lemma part_list_eq: + "(\X \ fst ` set xs. X) = UNIV + \ (\i < length xs. \j < length xs. i \ j + \ fst (xs ! i) \ fst (xs ! j) = {}) \ {} \ fst ` set xs + \ partition_on UNIV (set (map fst xs)) \ distinct (map fst xs)" + using part_list_set_eq_aux1 part_list_set_eq_aux2 + unfolding partition_on_def by (auto 5 0) + +text \@{typ 'd}: domain (such that the union of @{typ 'd} sets form a partition)\ +typedef ('d, 'a) part = "{xs :: ('d set \ 'a) list. partition_on UNIV (set (map fst xs)) \ distinct (map fst xs)}" + by (rule exI[of _ "[(UNIV, undefined)]"]) (auto simp: partition_on_def) + +setup_lifting type_definition_part + +lift_bnf (no_warn_wits, no_warn_transfer) (dead 'd, Vals: 'a) part + unfolding part_list_eq[symmetric] + by (auto simp: image_iff) + +subsection \\<^const>\size\ setup\ + +lift_definition subs :: "('d, 'a) part \ 'd set list" is "map fst" . + +lift_definition Subs :: "('d, 'a) part \ 'd set set" is "set o map fst" . + +lift_definition vals :: "('d, 'a) part \ 'a list" is "map snd" . + +lift_definition SubsVals :: "('d, 'a) part \ ('d set \ 'a) set" is "set" . + +lift_definition subsvals :: "('d, 'a) part \ ('d set \ 'a) list" is "id" . + +lift_definition size_part :: "('d \ nat) \ ('a \ nat) \ ('d, 'a) part \ nat" is "\f g. size_list (\(x, y). sum f x + g y)" . + +instantiation part :: (type, type) size begin + +definition size_part where +size_part_overloaded_def: "size_part = Partition.size_part (\_. 0) (\_. 0)" + +instance .. + +end + +lemma size_part_overloaded_simps[simp]: "size x = size (vals x)" + unfolding size_part_overloaded_def + by transfer (auto simp: size_list_conv_sum_list) + +lemma part_size_o_map: "inj h \ size_part f g \ map_part h = size_part f (g \ h)" + by (rule ext, transfer) + (auto simp: fun_eq_iff map_prod_def o_def split_beta case_prod_beta[abs_def]) + +setup \ +BNF_LFP_Size.register_size_global \<^type_name>\part\ \<^const_name>\size_part\ + @{thm size_part_overloaded_def} @{thms size_part_def size_part_overloaded_simps} + @{thms part_size_o_map} +\ + +lemma is_measure_size_part[measure_function]: "is_measure f \ is_measure g \ is_measure (size_part f g)" + by (rule is_measure_trivial) + +lemma size_part_estimation[termination_simp]: "x \ Vals xs \ y < g x \ y < size_part f g xs" + by transfer (auto simp: termination_simp) + +lemma size_part_estimation'[termination_simp]: "x \ Vals xs \ y \ g x \ y \ size_part f g xs" + by transfer (auto simp: termination_simp) + +lemma size_part_pointwise[termination_simp]: "(\x. x \ Vals xs \ f x \ g x) \ size_part h f xs \ size_part h g xs" + by transfer (force simp: image_iff intro!: size_list_pointwise) + +subsection \Functions on Valued Partitions\ + +lemma Vals_code[code]: "Vals x = set (map snd (Rep_part x))" + by (force simp: Vals_def) + +lemma Vals_transfer[transfer_rule]: "rel_fun (pcr_part (=) (=)) (=) (set \ map snd) Vals" + by (force simp: Vals_def rel_fun_def pcr_part_def cr_part_def rel_set_eq prod.rel_eq list.rel_eq) + +lemma set_vals[simp]: "set (vals xs) = Vals xs" + by transfer simp + +lift_definition part_hd :: "('d, 'a) part \ 'a" is "snd \ hd" . + +lift_definition tabulate :: "'d list \ ('d \ 'n) \ 'n \ ('d, 'n) part" is + "\ds f z. if distinct ds then if set ds = UNIV then map (\d. ({d}, f d)) ds else (- set ds, z) # map (\d. ({d}, f d)) ds else [(UNIV, z)]" + by (auto simp: o_def distinct_map inj_on_def partition_on_def disjoint_def) + +lift_definition lookup_part :: "('d, 'a) part \ 'd \ 'a" is "\xs d. snd (the (find (\(D, _). d \ D) xs))" . + +lemma Vals_tabulate[simp]: "Vals (tabulate xs f z) = + (if distinct xs then if set xs = UNIV then f ` set xs else {z} \ f ` set xs else {z})" + by transfer (auto simp: image_iff) + +lemma lookup_part_tabulate[simp]: "lookup_part (tabulate xs f z) x = + (if distinct xs \ x \ set xs then f x else z)" + by (transfer fixing: x xs f z) + (auto simp: find_dropWhile dropWhile_eq_Cons_conv map_eq_append_conv split: list.splits) + +lemma part_hd_Vals[simp]: "part_hd part \ Vals part" + by transfer (auto simp: partition_on_def image_iff intro!: hd_in_set) + +lemma lookup_part_Vals[simp]: "lookup_part part d \ Vals part" +proof (transfer, goal_cases part) + case (part xs d) + then have unique: "(\ij j \ fst (xs ! i) \ fst (xs ! j) = {})" + using part_list_eq[of xs] + by simp + from part obtain D x where D: "(D, x) \ set xs" "d \ D" + unfolding partition_on_def + by fastforce + with unique have "find (\(D, _). d \ D) xs = Some (D, x)" + unfolding set_eq_iff + by (auto simp: find_Some_iff in_set_conv_nth split_beta) + with D show ?case + by (force simp: image_iff) +qed + +lemma lookup_part_SubsVals: "\D. d \ D \ (D, lookup_part part d) \ SubsVals part" +proof (transfer, goal_cases part) + case (part d xs) + then have unique: "(\ij j \ fst (xs ! i) \ fst (xs ! j) = {})" + using part_list_eq[of xs] + by simp + from part obtain D x where D: "(D, x) \ set xs" "d \ D" + unfolding partition_on_def + by fastforce + with unique have "find (\(D, _). d \ D) xs = Some (D, x)" + unfolding set_eq_iff + by (auto simp: find_Some_iff in_set_conv_nth split_beta) + with D show ?case + by (force simp: image_iff) +qed + +lemma lookup_part_from_subvals: "(D, e) \ set (subsvals part) \ d \ D \ lookup_part part d = e" +proof (transfer fixing: D e d, goal_cases) + case (1 part) + then show ?case + proof (cases "find (\(D, _). d \ D) part") + case (Some De) + from 1 show ?thesis + unfolding partition_on_def set_eq_iff Some using Some unfolding find_Some_iff + by (fastforce dest!: spec[of _ d] simp: in_set_conv_nth split_beta dest: part_list_set_eq_aux2) + qed (auto simp: partition_on_def image_iff find_None_iff) +qed + +lemma size_lookup_part_estimation[termination_simp]: "size (lookup_part part d) < Suc (size_part (\_. 0) size part)" + unfolding less_Suc_eq_le + by (rule size_part_estimation'[OF _ order_refl]) simp + +lemma subsvals_part_estimation[termination_simp]: "(D, e) \ set (subsvals part) \ size e < Suc (size_part (\_. 0) size part)" + unfolding less_Suc_eq_le + by (rule size_part_estimation'[OF _ order_refl], transfer) + (force simp: image_iff) + +lemma size_part_hd_estimation[termination_simp]: "size (part_hd part) < Suc (size_part (\_. 0) size part)" + unfolding less_Suc_eq_le + by (rule size_part_estimation'[OF _ order_refl]) simp + +lemma size_last_estimation[termination_simp]: "xs \ [] \ size (last xs) < size_list size xs" + by (induct xs) auto + +lift_definition lookup :: "('d, 'a) part \ 'd \ ('d set \ 'a)" is "\xs d. the (find (\(D, _). d \ D) xs)" . + +lemma snd_lookup[simp]: "snd (lookup part d) = lookup_part part d" + by transfer auto + +lemma distinct_disjoint_uniq: "distinct xs \ disjoint (set xs) \ + i < j \ j < length xs \ d \ xs ! i \ d \ xs ! j \ False" + unfolding disjoint_def disjoint_iff + by (metis (no_types, lifting) order.strict_trans min.strict_order_iff nth_eq_iff_index_eq nth_mem) + +lemma partition_on_UNIV_find_Some: + "partition_on UNIV (set (map fst part)) \ distinct (map fst part) \ + \y. find (\(D, _). d \ D) part = Some y" + unfolding partition_on_def set_eq_iff + by (auto simp: find_Some_iff in_set_conv_nth + Ball_def image_iff Bex_def split_beta dest: distinct_disjoint_uniq dest!: spec[of _ d] + intro!: exI[where P="\x. \y z. P x y z \ Q x y z" for P Q, OF exI, OF exI, OF conjI]) + +lemma fst_lookup: "d \ fst (lookup part d)" +proof (transfer fixing: d, goal_cases) + case (1 part) + then obtain y where "find (\(D, _). d \ D) part = Some y" using partition_on_UNIV_find_Some + by fastforce + then show ?case + by (auto dest: find_Some_iff[THEN iffD1]) +qed + +lemma lookup_subsvals: "lookup part d \ set (subsvals part)" +proof (transfer fixing: d, goal_cases) + case (1 part) + then obtain y where "find (\(D, _). d \ D) part = Some y" using partition_on_UNIV_find_Some + by fastforce + then show ?case + by (auto simp: in_set_conv_nth dest: find_Some_iff[THEN iffD1]) +qed + +lift_definition trivial_part :: "'pt \ ('d, 'pt) part" is "\pt. [(UNIV, pt)]" + by (simp add: partition_on_space) + +lemma part_hd_trivial[simp]: "part_hd (trivial_part pt) = pt" + unfolding part_hd_def + by (transfer) simp + +lemma SubsVals_trivial[simp]: "SubsVals (trivial_part pt) = {(UNIV, pt)}" + unfolding SubsVals_def + by (transfer) simp + +section \Partitioned Decision Trees\ + +(* 'd: domain; 'l: leaf, 'n: variable *) +datatype (dead 'd, leaves: 'l, vars: 'n) pdt = Leaf (unleaf: 'l) | Node 'n "('d, ('d, 'l, 'n) pdt) part" + +inductive vars_order :: "'n list \ ('d, 'l, 'n) pdt \ bool" where + "vars_order vs (Leaf _)" +| "\expl \ Vals part1. vars_order vs expl \ vars_order (x # vs) (Node x part1)" +| "vars_order vs (Node x part1) \ x \ z \ vars_order (z # vs) (Node x part1)" + +lemma vars_order_Node: + assumes "distinct xs" + shows "vars_order xs (Node x part) = (\ys zs. xs = ys @ x # zs \ (\e \ Vals part. vars_order zs e))" +proof (safe, goal_cases LR RL) + case LR + then show ?case + by (induct xs "Node x part" rule: vars_order.induct) + (auto 4 3 intro: exI[of _ "_ # _"]) +next + case (RL ys zs) + with assms show ?case + by (induct ys arbitrary: xs) + (auto intro: vars_order.intros) +qed + +fun distinct_paths where + "distinct_paths (Leaf _) = True" +| "distinct_paths (Node x part) = (\e \ Vals part. x \ vars e \ distinct_paths e)" + +fun eval_pdt where + "eval_pdt v (Leaf l) = l" +| "eval_pdt v (Node x part) = eval_pdt v (lookup_part part (v x))" + +lemma eval_pdt_cong: "\x \ vars e. v x = v' x \ eval_pdt v e = eval_pdt v' e" + by (induct e) auto + +lemma vars_order_vars: "vars_order vs e \ vars e \ set vs" + by (induct vs e rule: vars_order.induct) auto + +lemma vars_order_distinct_paths: "vars_order vs e \ distinct vs \ distinct_paths e" + by (induct vs e rule: vars_order.induct) (auto dest!: vars_order_vars) + +lemma eval_pdt_fun_upd: "vars_order vs e \ x \ set vs \ eval_pdt (v(x := d)) e = eval_pdt v e" + by (induct vs e rule: vars_order.induct) auto + +(*<*) +end +(*>*) \ No newline at end of file diff --git a/thys/MFOTL_Checker/Prelim.thy b/thys/MFOTL_Checker/Prelim.thy new file mode 100644 --- /dev/null +++ b/thys/MFOTL_Checker/Prelim.thy @@ -0,0 +1,142 @@ +(*<*) +theory Prelim +imports Main +begin +(*>*) + +section \Auxiliary Lemmas\ + +lemma Cons_eq_upt_conv: "x # xs = [m ..< n] \ m < n \ x = m \ xs = [Suc m ..< n]" + by (induct n arbitrary: xs) (force simp: Cons_eq_append_conv)+ + +lemma map_setE[elim_format]: "map f xs = ys \ y \ set ys \ \x\set xs. f x = y" + by (induct xs arbitrary: ys) auto + +lemma set_Cons_eq: "set_Cons X XS = (\xs\XS. (\x. x # xs) ` X)" + by (auto simp: set_Cons_def) + +lemma set_Cons_empty_iff: "set_Cons X XS = {} \ (X = {} \ XS = {})" + by (auto simp: set_Cons_eq) + +lemma infinite_set_ConsI: + "XS \ {} \ infinite X \ infinite (set_Cons X XS)" + "X \ {} \ infinite XS \ infinite (set_Cons X XS)" +proof(unfold set_Cons_eq) + assume "infinite X" and "XS \ {}" + then obtain xs where "xs \ XS" + by blast + hence "inj (\x. x # xs)" + by (clarsimp simp: inj_on_def) + hence "infinite ((\x. x # xs) ` X)" + using \infinite X\ finite_imageD inj_on_def + by blast + moreover have "((\x. x # xs) ` X) \ (\xs\XS. (\x. x # xs) ` X)" + using \xs \ XS\ by auto + ultimately show "infinite (\xs\XS. (\x. x # xs) ` X)" + by (simp add: infinite_super) +next + assume "infinite XS" and "X \ {}" + then show "infinite (\xs\XS. (\x. x # xs) ` X)" + by (elim contrapos_nn finite_surj[of _ _ tl]) (auto simp: image_iff) +qed + +primrec fst_pos :: "'a list \ 'a \ nat option" + where "fst_pos [] x = None" + | "fst_pos (y#ys) x = (if x = y then Some 0 else + (case fst_pos ys x of None \ None | Some n \ Some (Suc n)))" + +lemma fst_pos_None_iff: "fst_pos xs x = None \ x \ set xs" + by (induct xs arbitrary: x; force split: option.splits) + +lemma nth_fst_pos: "x \ set xs \ xs ! (the (fst_pos xs x)) = x" + by (induct xs arbitrary: x; fastforce simp: fst_pos_None_iff split: option.splits) + +primrec positions :: "'a list \ 'a \ nat list" + where "positions [] x = []" + | "positions (y#ys) x = (\ns. if x = y then 0 # ns else ns) (map Suc (positions ys x))" + +lemma eq_positions_iff: "length xs = length ys + \ positions xs x = positions ys y \ (\n< length xs. xs ! n = x \ ys ! n = y)" + by (induct xs ys arbitrary: x y rule: list_induct2) (use less_Suc_eq_0_disj in auto) + +lemma positions_eq_nil_iff: "positions xs x = [] \ x \ set xs" + by (induct xs) simp_all + +lemma positions_nth: "n \ set (positions xs x) \ xs ! n = x" + by (induct xs arbitrary: n x) + (auto simp: positions_eq_nil_iff[symmetric] split: if_splits) + +lemma set_positions_eq: "set (positions xs x) = {n. xs ! n = x \ n < length xs}" + by (induct xs arbitrary: x) + (use less_Suc_eq_0_disj in \auto simp: positions_eq_nil_iff[symmetric] image_iff split: if_splits\) + +lemma positions_length: "n \ set (positions xs x) \ n < length xs" + by (induct xs arbitrary: n x) + (auto simp: positions_eq_nil_iff[symmetric] split: if_splits) + +lemma positions_nth_cong: + "m \ set (positions xs x) \ n \ set (positions xs x) \ xs ! n = xs ! m" + using positions_nth[of _ xs x] by simp + +lemma fst_pos_in_positions: "x \ set xs \ the (fst_pos xs x) \ set (positions xs x)" + by (induct xs arbitrary: x, simp) + (fastforce simp: hd_map fst_pos_None_iff split: option.splits) + +lemma hd_positions_eq_fst_pos: "x \ set xs \ hd (positions xs x) = the (fst_pos xs x)" + by (induct xs arbitrary: x) + (auto simp: hd_map fst_pos_None_iff positions_eq_nil_iff split: option.splits) + +lemma sorted_positions: "sorted (positions xs x)" + by (induct xs arbitrary: x) (auto simp add: sorted_iff_nth_Suc nth_Cons' gr0_conv_Suc) + +lemma Min_sorted_list: "sorted xs \ xs \ [] \ Min (set xs) = hd xs" + by (induct xs) + (auto simp: Min_insert2) + +lemma Min_positions: "x \ set xs \ Min (set (positions xs x)) = the (fst_pos xs x)" + by (auto simp: Min_sorted_list[OF sorted_positions] + positions_eq_nil_iff hd_positions_eq_fst_pos) + +lemma subset_positions_map_fst: "set (positions tXs tX) \ set (positions (map fst tXs) (fst tX))" + by (induct tXs arbitrary: tX) + (auto simp: subset_eq) + +lemma subset_positions_map_snd: "set (positions tXs tX) \ set (positions (map snd tXs) (snd tX))" + by (induct tXs arbitrary: tX) + (auto simp: subset_eq) + +lemma Max_eqI: "finite A \ A \ {} \ (\a. a \ A \ a \ b) \ \a\A. b \ a \ Max A = b" + by (rule antisym[OF Max.boundedI Max_ge_iff[THEN iffD2]]; clarsimp) + +lemma Max_Suc: "X \ {} \ finite X \ Max (Suc ` X) = Suc (Max X)" + using Max_ge Max_in + by (intro Max_eqI) blast+ + +lemma Max_insert0: "X \ {} \ finite X \ Max (insert (0::nat) X) = Max X" + using Max_ge Max_in + by (intro Max_eqI) blast+ + +lemma positions_Cons_notin_tail: "x \ set xs \ positions (x # xs) x = [0::nat]" + by (cases xs) (auto simp: positions_eq_nil_iff) + +lemma Max_set_positions_Cons_hd: + "x \ set xs \ Max (set (positions (x # xs) x)) = 0" + by (subst positions_Cons_notin_tail) simp_all + +lemma Max_set_positions_Cons_tl: + "y \ set xs \ Max (set (positions (x # xs) y)) = Suc (Max (set (positions xs y)))" + by (auto simp: Max_Suc positions_eq_nil_iff) + +lemma max_aux: "finite X \ Suc j \ X \ Max (insert (Suc j) (X - {j})) = Max (insert j X)" + by (smt (verit) max.orderI Max.insert_remove Max_ge Max_insert empty_iff insert_Diff_single + insert_absorb insert_iff max_def not_less_eq_eq) + +lemma ball_swap: "(\x \ A. \y \ B. P x y) = (\y \ B. \x \ A. P x y)" + by auto + +lemma ball_triv_nonempty: "A \ {} \ (\x \ A. P) = P" + by auto + +(*<*) +end +(*>*) \ No newline at end of file diff --git a/thys/MFOTL_Checker/Proof_Object.thy b/thys/MFOTL_Checker/Proof_Object.thy new file mode 100644 --- /dev/null +++ b/thys/MFOTL_Checker/Proof_Object.thy @@ -0,0 +1,123 @@ +(*<*) +theory Proof_Object + imports Formula Partition +begin +(*>*) + +section \Proof Objects\ + +datatype (dead 'n, dead 'd) sproof = STT nat + | SPred nat 'n "('n, 'd) Formula.trm list" + | SEq_Const nat 'n 'd + | SNeg "('n, 'd) vproof" + | SOrL "('n, 'd) sproof" + | SOrR "('n, 'd) sproof" + | SAnd "('n, 'd) sproof" "('n, 'd) sproof" + | SImpL "('n, 'd) vproof" + | SImpR "('n, 'd) sproof" + | SIffSS "('n, 'd) sproof" "('n, 'd) sproof" + | SIffVV "('n, 'd) vproof" "('n, 'd) vproof" + | SExists 'n 'd "('n, 'd) sproof" + | SForall 'n "('d, ('n, 'd) sproof) part" + | SPrev "('n, 'd) sproof" + | SNext "('n, 'd) sproof" + | SOnce nat "('n, 'd) sproof" + | SEventually nat "('n, 'd) sproof" + | SHistorically nat nat "('n, 'd) sproof list" + | SHistoricallyOut nat + | SAlways nat nat "('n, 'd) sproof list" + | SSince "('n, 'd) sproof" "('n, 'd) sproof list" + | SUntil "('n, 'd) sproof list" "('n, 'd) sproof" + and ('n, 'd) vproof = VFF nat + | VPred nat 'n "('n, 'd) Formula.trm list" + | VEq_Const nat 'n 'd + | VNeg "('n, 'd) sproof" + | VOr "('n, 'd) vproof" "('n, 'd) vproof" + | VAndL "('n, 'd) vproof" + | VAndR "('n, 'd) vproof" + | VImp "('n, 'd) sproof" "('n, 'd) vproof" + | VIffSV "('n, 'd) sproof" "('n, 'd) vproof" + | VIffVS "('n, 'd) vproof" "('n, 'd) sproof" + | VExists 'n "('d, ('n, 'd) vproof) part" + | VForall 'n 'd "('n, 'd) vproof" + | VPrev "('n, 'd) vproof" + | VPrevZ + | VPrevOutL nat + | VPrevOutR nat + | VNext "('n, 'd) vproof" + | VNextOutL nat + | VNextOutR nat + | VOnceOut nat + | VOnce nat nat "('n, 'd) vproof list" + | VEventually nat nat "('n, 'd) vproof list" + | VHistorically nat "('n, 'd) vproof" + | VAlways nat "('n, 'd) vproof" + | VSinceOut nat + | VSince nat "('n, 'd) vproof" "('n, 'd) vproof list" + | VSinceInf nat nat "('n, 'd) vproof list" + | VUntil nat "('n, 'd) vproof list" "('n, 'd) vproof" + | VUntilInf nat nat "('n, 'd) vproof list" + +type_synonym ('n, 'd) "proof" = "('n, 'd) sproof + ('n, 'd) vproof" + +type_synonym ('n, 'd) expl = "('d, ('n, 'd) proof, 'n) pdt" + +fun s_at :: "('n, 'd) sproof \ nat" and + v_at :: "('n, 'd) vproof \ nat" where + "s_at (STT i) = i" +| "s_at (SPred i _ _) = i" +| "s_at (SEq_Const i _ _) = i" +| "s_at (SNeg vp) = v_at vp" +| "s_at (SOrL sp1) = s_at sp1" +| "s_at (SOrR sp2) = s_at sp2" +| "s_at (SAnd sp1 _) = s_at sp1" +| "s_at (SImpL vp1) = v_at vp1" +| "s_at (SImpR sp2) = s_at sp2" +| "s_at (SIffSS sp1 _) = s_at sp1" +| "s_at (SIffVV vp1 _) = v_at vp1" +| "s_at (SExists _ _ sp) = s_at sp" +| "s_at (SForall _ part) = s_at (part_hd part)" +| "s_at (SPrev sp) = s_at sp + 1" +| "s_at (SNext sp) = s_at sp - 1" +| "s_at (SOnce i _) = i" +| "s_at (SEventually i _) = i" +| "s_at (SHistorically i _ _) = i" +| "s_at (SHistoricallyOut i) = i" +| "s_at (SAlways i _ _) = i" +| "s_at (SSince sp2 sp1s) = (case sp1s of [] \ s_at sp2 | _ \ s_at (last sp1s))" +| "s_at (SUntil sp1s sp2) = (case sp1s of [] \ s_at sp2 | sp1 # _ \ s_at sp1)" +| "v_at (VFF i) = i" +| "v_at (VPred i _ _) = i" +| "v_at (VEq_Const i _ _) = i" +| "v_at (VNeg sp) = s_at sp" +| "v_at (VOr vp1 _) = v_at vp1" +| "v_at (VAndL vp1) = v_at vp1" +| "v_at (VAndR vp2) = v_at vp2" +| "v_at (VImp sp1 _) = s_at sp1" +| "v_at (VIffSV sp1 _) = s_at sp1" +| "v_at (VIffVS vp1 _) = v_at vp1" +| "v_at (VExists _ part) = v_at (part_hd part)" +| "v_at (VForall _ _ vp1) = v_at vp1" +| "v_at (VPrev vp) = v_at vp + 1" +| "v_at (VPrevZ) = 0" +| "v_at (VPrevOutL i) = i" +| "v_at (VPrevOutR i) = i" +| "v_at (VNext vp) = v_at vp - 1" +| "v_at (VNextOutL i) = i" +| "v_at (VNextOutR i) = i" +| "v_at (VOnceOut i) = i" +| "v_at (VOnce i _ _) = i" +| "v_at (VEventually i _ _) = i" +| "v_at (VHistorically i _) = i" +| "v_at (VAlways i _) = i" +| "v_at (VSinceOut i) = i" +| "v_at (VSince i _ _) = i" +| "v_at (VSinceInf i _ _) = i" +| "v_at (VUntil i _ _) = i" +| "v_at (VUntilInf i _ _) = i" + +definition p_at :: "('n, 'd) proof \ nat" where "p_at p = case_sum s_at v_at p" + +(*<*) +end +(*>*) diff --git a/thys/MFOTL_Checker/Proof_System.thy b/thys/MFOTL_Checker/Proof_System.thy new file mode 100644 --- /dev/null +++ b/thys/MFOTL_Checker/Proof_System.thy @@ -0,0 +1,534 @@ +(*<*) +theory Proof_System + imports Formula Partition +begin +(*>*) + +section \Proof System\ + +unbundle MFOTL_notation + +context begin + +inductive SAT and VIO :: "('n, 'd) trace \ ('n, 'd) env \ nat \ ('n, 'd) formula \ bool" for \ where + STT: "SAT \ v i TT" +| VFF: "VIO \ v i FF" +| SPred: "(r, eval_trms v ts) \ \ \ i \ SAT \ v i (Pred r ts)" +| VPred: "(r, eval_trms v ts) \ \ \ i \ VIO \ v i (Pred r ts)" +| SEq_Const: "v x = c \ SAT \ v i (Eq_Const x c)" +| VEq_Const: "v x \ c \ VIO \ v i (Eq_Const x c)" +| SNeg: "VIO \ v i \ \ SAT \ v i (Neg \)" +| VNeg: "SAT \ v i \ \ VIO \ v i (Neg \)" +| SOrL: "SAT \ v i \ \ SAT \ v i (Or \ \)" +| SOrR: "SAT \ v i \ \ SAT \ v i (Or \ \)" +| VOr: "VIO \ v i \ \ VIO \ v i \ \ VIO \ v i (Or \ \)" +| SAnd: "SAT \ v i \ \ SAT \ v i \ \ SAT \ v i (And \ \)" +| VAndL: "VIO \ v i \ \ VIO \ v i (And \ \)" +| VAndR: "VIO \ v i \ \ VIO \ v i (And \ \)" +| SImpL: "VIO \ v i \ \ SAT \ v i (Imp \ \)" +| SImpR: "SAT \ v i \ \ SAT \ v i (Imp \ \)" +| VImp: "SAT \ v i \ \ VIO \ v i \ \ VIO \ v i (Imp \ \)" +| SIffSS: "SAT \ v i \ \ SAT \ v i \ \ SAT \ v i (Iff \ \)" +| SIffVV: "VIO \ v i \ \ VIO \ v i \ \ SAT \ v i (Iff \ \)" +| VIffSV: "SAT \ v i \ \ VIO \ v i \ \ VIO \ v i (Iff \ \)" +| VIffVS: "VIO \ v i \ \ SAT \ v i \ \ VIO \ v i (Iff \ \)" +| SExists: "\z. SAT \ (v (x := z)) i \ \ SAT \ v i (Exists x \)" +| VExists: "\z. VIO \ (v (x := z)) i \ \ VIO \ v i (Exists x \)" +| SForall: "\z. SAT \ (v (x := z)) i \ \ SAT \ v i (Forall x \)" +| VForall: "\z. VIO \ (v (x := z)) i \ \ VIO \ v i (Forall x \)" +| SPrev: "i > 0 \ mem (\ \ i) I \ SAT \ v (i-1) \ \ SAT \ v i (\<^bold>Y I \)" +| VPrev: "i > 0 \ VIO \ v (i-1) \ \ VIO \ v i (\<^bold>Y I \)" +| VPrevZ: "i = 0 \ VIO \ v i (\<^bold>Y I \)" +| VPrevOutL: "i > 0 \ (\ \ i) < (left I) \ VIO \ v i (\<^bold>Y I \)" +| VPrevOutR: "i > 0 \ enat (\ \ i) > (right I) \ VIO \ v i (\<^bold>Y I \)" +| SNext: "mem (\ \ (i+1)) I \ SAT \ v (i+1) \ \ SAT \ v i (\<^bold>X I \)" +| VNext: "VIO \ v (i+1) \ \ VIO \ v i (\<^bold>X I \)" +| VNextOutL: "(\ \ (i+1)) < (left I) \ VIO \ v i (\<^bold>X I \)" +| VNextOutR: "enat (\ \ (i+1)) > (right I) \ VIO \ v i (\<^bold>X I \)" +| SOnce: "j \ i \ mem (\ \ i j) I \ SAT \ v j \ \ SAT \ v i (\<^bold>P I \)" +| VOnceOut: "\ \ i < \ \ 0 + left I \ VIO \ v i (\<^bold>P I \)" +| VOnce: "j = (case right I of \ \ 0 + | enat b \ ETP_p \ i b) \ + (\ \ i) \ (\ \ 0) + left I \ + (\k. k \ {j .. LTP_p \ i I} \ VIO \ v k \) \ VIO \ v i (\<^bold>P I \)" +| SEventually: "j \ i \ mem (\ \ j i) I \ SAT \ v j \ \ SAT \ v i (\<^bold>F I \)" +| VEventually: "(\k. k \ (case right I of \ \ {ETP_f \ i I ..} + | enat b \ {ETP_f \ i I .. LTP_f \ i b}) \ VIO \ v k \) \ + VIO \ v i (\<^bold>F I \)" +| SHistorically: "j = (case right I of \ \ 0 + | enat b \ ETP_p \ i b) \ + (\ \ i) \ (\ \ 0) + left I \ + (\k. k \ {j .. LTP_p \ i I} \ SAT \ v k \) \ SAT \ v i (\<^bold>H I \)" +| SHistoricallyOut: "\ \ i < \ \ 0 + left I \ SAT \ v i (\<^bold>H I \)" +| VHistorically: "j \ i \ mem (\ \ i j) I \ VIO \ v j \ \ VIO \ v i (\<^bold>H I \)" +| SAlways: "(\k. k \ (case right I of \ \ {ETP_f \ i I ..} + | enat b \ {ETP_f \ i I .. LTP_f \ i b}) \ SAT \ v k \) \ + SAT \ v i (\<^bold>G I \)" +| VAlways: "j \ i \ mem (\ \ j i) I \ VIO \ v j \ \ VIO \ v i (\<^bold>G I \)" +| SSince: "j \ i \ mem (\ \ i j) I \ SAT \ v j \ \ (\k. k \ {j <.. i} \ + SAT \ v k \) \ SAT \ v i (\ \<^bold>S I \)" +| VSinceOut: "\ \ i < \ \ 0 + left I \ VIO \ v i (\ \<^bold>S I \)" +| VSince: "(case right I of \ \ True + | enat b \ ETP \ ((\ \ i) - b) \ j) \ + j \ i \ (\ \ 0) + left I \ (\ \ i) \ VIO \ v j \ \ + (\k. k \ {j .. (LTP_p \ i I)} \ VIO \ v k \) \ VIO \ v i (\ \<^bold>S I \)" +| VSinceInf: "j = (case right I of \ \ 0 + | enat b \ ETP_p \ i b) \ + (\ \ i) \ (\ \ 0) + left I \ + (\k. k \ {j .. LTP_p \ i I} \ VIO \ v k \) \ VIO \ v i (\ \<^bold>S I \)" +| SUntil: "j \ i \ mem (\ \ j i) I \ SAT \ v j \ \ (\k. k \ {i ..< j} \ SAT \ v k \) \ + SAT \ v i (\ \<^bold>U I \)" +| VUntil: "(case right I of \ \ True + | enat b \ j < LTP_f \ i b) \ + j \ i \ VIO \ v j \ \ (\k. k \ {ETP_f \ i I .. j} \ VIO \ v k \) \ + VIO \ v i (\ \<^bold>U I \)" +| VUntilInf: "(\k. k \ (case right I of \ \ {ETP_f \ i I ..} + | enat b \ {ETP_f \ i I .. LTP_f \ i b}) \ VIO \ v k \) \ + VIO \ v i (\ \<^bold>U I \)" + +subsection \Soundness and Completeness\ + +lemma not_sat_SinceD: + assumes unsat: "\ \\, v, i\ \ \ \<^bold>S I \" and + witness: "\j \ i. mem (\ \ i - \ \ j) I \ \\, v, j\ \ \" + shows "\j \ i. ETP \ (case right I of \ \ 0 | enat n \ \ \ i - n) \ j \ \ \\, v, j\ \ \ + \ (\k \ {j .. (min i (LTP \ (\ \ i - left I)))}. \ \\, v, k\ \ \)" +proof - + define A and j where A_def: "A \ {j. j \ i \ mem (\ \ i - \ \ j) I \ \\, v, j\ \ \}" + and j_def: "j \ Max A" + from witness have j: "j \ i" "\\, v, j\ \ \" "mem (\ \ i - \ \ j) I" + using Max_in[of A] unfolding j_def[symmetric] unfolding A_def + by auto + moreover + from j(3) have "ETP \ (case right I of enat n \ \ \ i - n | \ \ 0) \ j" + unfolding ETP_def by (intro Least_le) (auto split: enat.splits) + moreover + { fix j + assume j: "\ \ j \ \ \ i" + then obtain k where k: "\ \ i < \ \ k" + by (meson ex_le_\ gt_ex less_le_trans) + have "j \ ETP \ (Suc (\ \ i))" + unfolding ETP_def + proof (intro LeastI2[of _ k "\i. j \ i"]) + fix l + assume "Suc (\ \ i) \ \ \ l" + with j show "j \ l" + by (metis lessI less_\D nless_le order_less_le_trans) + qed (auto simp: Suc_le_eq k(1)) + } note * = this + { fix k + assume "k \ {j <.. (min i (LTP \ (\ \ i - left I)))}" + with j(3) have k: "j < k" "k \ i" "k \ Max {j. left I + \ \ j \ \ \ i}" + by (auto simp: LTP_def le_diff_conv2 add.commute) + with j(3) obtain l where "left I + \ \ l \ \ \ i" "k \ l" + by (subst (asm) Max_ge_iff) (auto simp: le_diff_conv2 * + intro!: finite_subset[of _ "{0 .. ETP \ (\ \ i + 1)}"]) + then have "mem (\ \ i - \ \ k) I" + using k(1,2) j(3) + by (cases "right I") (auto simp: le_diff_conv le_diff_conv2 add.commute dest: \_mono + elim: order_trans[rotated] order_trans) + with Max_ge[of A k] k have "\ \\, v, k\ \ \" + unfolding j_def[symmetric] unfolding A_def + by auto + } + ultimately show ?thesis using unsat + by (auto dest!: spec[of _ j]) +qed + +lemma not_sat_UntilD: + assumes unsat: "\ \\, v, i\ \ \ \<^bold>U I \" + and witness: "\j \ i. mem (\ \ j i) I \ \\, v, j\ \ \" + shows "\j \ i. (case right I of \ \ True | enat n \ j < LTP \ (\ \ i + n)) + \ \ (\\, v, j\ \ \) \ (\k \ {(max i (ETP \ (\ \ i + left I))) .. j}. + \ \\, v, k\ \ \)" +proof - + from \_mono have i0: "\ \ 0 \ \ \ i" by auto + from witness obtain jmax where jmax: "jmax \ i" "\\, v, jmax\ \ \" + "mem (\ \ jmax i) I" by blast + define A and j where A_def: "A \ {j. j \ i \ j \ jmax + \ mem (\ \ j i) I \ \\, v, j\ \ \}" and j_def: "j \ Min A" + have j: "j \ i" "\\, v, j\ \ \" "mem (\ \ j i) I" + using A_def j_def jmax Min_in[of A] + unfolding j_def[symmetric] unfolding A_def + by fastforce+ + moreover have "case right I of \ \ True | enat n \ j \ LTP \ (\ \ i + n)" + using i0 j(1,3) + by (auto simp: i_LTP_tau trans_le_add1 split: enat.splits) + moreover + { fix k + assume k_def: "k \ {(max i (ETP \ (\ \ i + left I))) ..< j}" + then have ki: "\ \ k \ \ \ i + left I" using i_ETP_tau by auto + with k_def have kj: "k < j" by auto + then have "\ \ k \ \ \ j" by auto + then have "\ \ k i \ \ \ j i" by auto + with this j(3) have "enat (\ \ k i) \ right I" + by (meson enat_ord_simps(1) order_subst2) + with this ki j(3) have mem_k: "mem (\ \ k i) I" + unfolding ETP_def by (auto simp: Least_le) + + with j_def have "j \ jmax" using Min_in[of A] + using jmax A_def + by (metis (mono_tags, lifting) Collect_empty_eq + finite_nat_set_iff_bounded_le mem_Collect_eq order_refl) + with this k_def have kjm: "k \ jmax" by auto + + with this mem_k ki Min_le[of A k] k_def have "k \ A" + unfolding j_def[symmetric] unfolding A_def unfolding ETP_def + using finite_nat_set_iff_bounded_le kj leD by blast + with this mem_k k_def kjm have "\ \\, v, k\ \ \" + by (simp add: A_def) } + ultimately show ?thesis using unsat + by (auto split: enat.splits dest!: spec[of _ j]) +qed + +lemma soundness_raw: "(SAT \ v i \ \ \\, v, i\ \ \) \ (VIO \ v i \ \ \ \\, v, i\ \ \)" +proof (induct v i \ rule: SAT_VIO.induct) + case (VOnceOut i I v \) + { fix j + from \_mono have j0: "\ \ 0 \ \ \ j" by auto + then have "\ \ i < \ \ j + left I" using VOnceOut by linarith + then have "\ \ i j < left I" + using VOnceOut less_\D verit_comp_simplify1(3) by fastforce + then have "\ mem (\ \ i j) I" by auto } + then show ?case + by auto +next + case (VOnce j I i v \) + { fix k + assume k_def: "\\, v, k\ \ \ \ mem (\ \ i k) I \ k \ i" + then have k_tau: "\ \ k \ \ \ i - left I" + using diff_le_mono2 by fastforce + then have k_ltp: "k \ LTP \ (\ \ i - left I)" + using VOnce i_LTP_tau add_le_imp_le_diff + by blast + then have "k \ {j .. LTP_p \ i I}" + using k_def VOnce k_tau + by auto + then have "k < j" using k_def k_ltp by auto } + then show ?case + using VOnce + by (cases "right I = \") + (auto 0 0 simp: i_ETP_tau i_LTP_tau le_diff_conv2) +next + case (VEventually I i v \) + { fix k n + assume r: "right I = enat n" + from this have tin0: "\ \ i + n \ \ \ 0" + by (auto simp add: trans_le_add1) + define j where "j = LTP \ ((\ \ i) + n)" + then have j_i: "i \ j" + by (auto simp add: i_LTP_tau trans_le_add1 j_def) + assume k_def: "\\, v, k\ \ \ \ mem (\ \ k i) I \ i \ k" + then have "\ \ k \ \ \ i + left I" + using le_diff_conv2 by auto + then have k_etp: "k \ ETP \ (\ \ i + left I)" + using i_ETP_tau by blast + from this k_def VEventually have "k \ {ETP_f \ i I .. j}" + by (auto simp: r j_def) + then have "j < k" using r k_def k_etp by auto + from k_def r have "\ \ k i \ n" by auto + then have "\ \ k \ \ \ i + n" by auto + then have "k \ j" using tin0 i_LTP_tau by (auto simp add: j_def) } + note aux = this + show ?case + proof (cases "right I") + case (enat n) + show ?thesis + using VEventually[unfolded enat, simplified] aux + by (simp add: i_ETP_tau enat) + (metis \_mono le_add_diff_inverse nat_add_left_cancel_le) + next + case infinity + show ?thesis + using VEventually + by (auto simp: infinity i_ETP_tau le_diff_conv2) + qed +next + case (SHistorically j I i v \) + { fix k + assume k_def: "\ \\, v, k\ \ \ \ mem (\ \ i k) I \ k \ i" + then have k_tau: "\ \ k \ \ \ i - left I" + using diff_le_mono2 by fastforce + then have k_ltp: "k \ LTP \ (\ \ i - left I)" + using SHistorically i_LTP_tau add_le_imp_le_diff + by blast + then have "k \ {j .. LTP_p \ i I}" + using k_def SHistorically k_tau + by auto + then have "k < j" using k_def k_ltp by auto } + then show ?case + using SHistorically + by (cases "right I = \") + (auto 0 0 simp add: le_diff_conv2 i_ETP_tau i_LTP_tau) +next + case (SHistoricallyOut i I v \) + { fix j + from \_mono have j0: "\ \ 0 \ \ \ j" by auto + then have "\ \ i < \ \ j + left I" using SHistoricallyOut by linarith + then have "\ \ i j < left I" + using SHistoricallyOut less_\D not_le by fastforce + then have "\ mem (\ \ i j) I" by auto } + then show ?case by auto +next + case (SAlways I i v \) + { fix k n + assume r: "right I = enat n" + from this SAlways have tin0: "\ \ i + n \ \ \ 0" + by (auto simp add: trans_le_add1) + define j where "j = LTP \ ((\ \ i) + n)" + from SAlways have j_i: "i \ j" + by (auto simp add: i_LTP_tau trans_le_add1 j_def) + assume k_def: "\ \\, v, k\ \ \ \ mem (\ \ k i) I \ i \ k" + then have "\ \ k \ \ \ i + left I" + using le_diff_conv2 by auto + then have k_etp: "k \ ETP \ (\ \ i + left I)" + using SAlways i_ETP_tau by blast + from this k_def SAlways have "k \ {ETP_f \ i I .. j}" + by (auto simp: r j_def) + then have "j < k" using SAlways k_def k_etp by simp + from k_def r have "\ \ k i \ n" by simp + then have "\ \ k \ \ \ i + n" by simp + then have "k \ j" + using tin0 i_LTP_tau + by (auto simp add: j_def) } + note aux = this + show ?case + proof (cases "right I") + case (enat n) + show ?thesis + using SAlways[unfolded enat, simplified] aux + by (simp add: i_ETP_tau le_diff_conv2 enat) + (metis Groups.ab_semigroup_add_class.add.commute add_le_imp_le_diff) + next + case infinity + show ?thesis + using SAlways + by (auto simp: infinity i_ETP_tau le_diff_conv2) + qed +next + case (VSinceOut i I v \ \) + { fix j + from \_mono have j0: "\ \ 0 \ \ \ j" by auto + then have "\ \ i < \ \ j + left I" using VSinceOut by linarith + then have "\ \ i j < left I" using VSinceOut j0 + by (metis add.commute gr_zeroI leI less_\D less_diff_conv2 order_less_imp_not_less zero_less_diff) + then have "\ mem (\ \ i j) I" by auto } + then show ?case using VSinceOut by auto +next + case (VSince I i j v \ \) + { fix k + assume k_def: "\\, v, k\ \ \ \ mem (\ \ i k) I \ k \ i" + then have "\ \ k \ \ \ i - left I" using diff_le_mono2 by fastforce + then have k_ltp: "k \ LTP \ (\ \ i - left I)" + using VSince i_LTP_tau add_le_imp_le_diff + by blast + then have "k < j" using k_def VSince(7)[of k] + by force + then have "j \ {k <.. i} \ \ \\, v, j\ \ \" using VSince + by auto } + then show ?case using VSince + by force +next + case (VSinceInf j I i v \ \) + { fix k + assume k_def: "\\, v, k\ \ \ \ mem (\ \ i k) I \ k \ i" + then have k_tau: "\ \ k \ \ \ i - left I" + using diff_le_mono2 by fastforce + then have k_ltp: "k \ LTP \ (\ \ i - left I)" + using VSinceInf i_LTP_tau add_le_imp_le_diff + by blast + then have "k \ {j .. LTP_p \ i I}" + using k_def VSinceInf k_tau + by auto + then have "k < j" using k_def k_ltp by auto } + then show ?case + using VSinceInf + by (cases "right I = \") + (auto 0 0 simp: i_ETP_tau i_LTP_tau le_diff_conv2) +next + case (VUntil I j i v \ \) + { fix k + assume k_def: "\\, v, k\ \ \ \ mem (\ \ k i) I \ i \ k" + then have "\ \ k \ \ \ i + left I" + using le_diff_conv2 by auto + then have k_etp: "k \ ETP \ (\ \ i + left I)" + using VUntil i_ETP_tau by blast + from this k_def VUntil have "k \ {ETP_f \ i I .. j}" by auto + then have "j < k" using k_etp k_def by auto + then have "j \ {i ..< k} \ VIO \ v j \" using VUntil k_def + by auto } + then show ?case + using VUntil by force +next + case (VUntilInf I i v \ \) + { fix k n + assume r: "right I = enat n" + from this VUntilInf have tin0: "\ \ i + n \ \ \ 0" + by (auto simp add: trans_le_add1) + define j where "j = LTP \ ((\ \ i) + n)" + from VUntilInf have j_i: "i \ j" + by (auto simp add: i_LTP_tau trans_le_add1 j_def) + assume k_def: "\\, v, k\ \ \ \ mem (\ \ k i) I \ i \ k" + then have "\ \ k \ \ \ i + left I" + using le_diff_conv2 by auto + then have k_etp: "k \ ETP \ (\ \ i + left I)" + using VUntilInf i_ETP_tau by blast + from this k_def VUntilInf have "k \ {ETP_f \ i I .. j}" + by (auto simp: r j_def) + then have "j < k" using VUntilInf k_def k_etp by auto + from k_def r have "\ \ k i \ n" by auto + then have "\ \ k \ \ \ i + n" by auto + then have "k \ j" + using tin0 VUntilInf i_LTP_tau r k_def + by (force simp add: j_def) } + note aux = this + show ?case + proof (cases "right I") + case (enat n) + show ?thesis + using VUntilInf[unfolded enat, simplified] aux + by (simp add: i_ETP_tau enat) + (metis \_mono le_add_diff_inverse nat_add_left_cancel_le) + next + case infinity + show ?thesis + using VUntilInf + by (auto simp: infinity i_ETP_tau le_diff_conv2) + qed +qed (auto simp: fun_upd_def split: nat.splits) + +lemmas soundness = soundness_raw[THEN conjunct1, THEN mp] soundness_raw[THEN conjunct2, THEN mp] + +lemma completeness_raw: "(\\, v, i\ \ \ \ SAT \ v i \) \ (\ \\, v, i\ \ \ \ VIO \ v i \)" +proof (induct \ arbitrary: v i) + case (Prev I \) + show ?case using Prev + by (auto intro: SAT_VIO.SPrev SAT_VIO.VPrev SAT_VIO.VPrevOutL SAT_VIO.VPrevOutR SAT_VIO.VPrevZ split: nat.splits) +next + case (Once I \) + { assume "\\, v, i\ \ \<^bold>P I \" + with Once have "SAT \ v i (\<^bold>P I \)" + by (auto intro: SAT_VIO.SOnce) } + moreover + { assume i_l: "\ \ i < \ \ 0 + left I" + with Once have "VIO \ v i (\<^bold>P I \)" + by (auto intro: SAT_VIO.VOnceOut) } + moreover + { assume unsat: "\ \\, v, i\ \ \<^bold>P I \" + and i_ge: "\ \ 0 + left I \ \ \ i" + with Once have "VIO \ v i (\<^bold>P I \)" + by (auto intro!: SAT_VIO.VOnce simp: i_LTP_tau i_ETP_tau + split: enat.splits) } + ultimately show ?case + by force +next + case (Historically I \) + from \_mono have i0: "\ \ 0 \ \ \ i" by auto + { assume sat: "\\, v, i\ \ \<^bold>H I \" + and i_ge: "\ \ i \ \ \ 0 + left I" + with Historically have "SAT \ v i (\<^bold>H I \)" + using le_diff_conv + by (auto intro!: SAT_VIO.SHistorically simp: i_LTP_tau i_ETP_tau + split: enat.splits) } + moreover + { assume "\ \\, v, i\ \ \<^bold>H I \" + with Historically have "VIO \ v i (\<^bold>H I \)" + by (auto intro: SAT_VIO.VHistorically) } + moreover + { assume i_l: "\ \ i < \ \ 0 + left I" + with Historically have "SAT \ v i (\<^bold>H I \)" + by (auto intro: SAT_VIO.SHistoricallyOut) } + ultimately show ?case + by force +next + case (Eventually I \) + from \_mono have i0: "\ \ 0 \ \ \ i" by auto + { assume "\\, v, i\ \ \<^bold>F I \" + with Eventually have "SAT \ v i (\<^bold>F I \)" + by (auto intro: SAT_VIO.SEventually) } + moreover + { assume unsat: "\ \\, v, i\ \ \<^bold>F I \" + with Eventually have "VIO \ v i (\<^bold>F I \)" + by (auto intro!: SAT_VIO.VEventually simp: add_increasing2 i0 i_LTP_tau i_ETP_tau + split: enat.splits) } + ultimately show ?case by auto +next + case (Always I \) + from \_mono have i0: "\ \ 0 \ \ \ i" by auto + { assume "\ \\, v, i\ \ \<^bold>G I \" + with Always have "VIO \ v i (\<^bold>G I \)" + by (auto intro: SAT_VIO.VAlways) } + moreover + { assume sat: "\\, v, i\ \ \<^bold>G I \" + with Always have "SAT \ v i (\<^bold>G I \)" + by (auto intro!: SAT_VIO.SAlways simp: add_increasing2 i0 i_LTP_tau i_ETP_tau le_diff_conv split: enat.splits)} + ultimately show ?case by auto +next + case (Since \ I \) + { assume "\\, v, i\ \ \ \<^bold>S I \" + with Since have "SAT \ v i (\ \<^bold>S I \)" + by (auto intro: SAT_VIO.SSince) } + moreover + { assume i_l: "\ \ i < \ \ 0 + left I" + with Since have "VIO \ v i (\ \<^bold>S I \)" + by (auto intro: SAT_VIO.VSinceOut) } + moreover + { assume unsat: "\ \\, v, i\ \ \ \<^bold>S I \" + and nw: "\j\i. \ mem (\ \ i j) I \ \ \\, v, j\ \ \" + and i_ge: "\ \ 0 + left I \ \ \ i" + with Since have "VIO \ v i (\ \<^bold>S I \)" + by (auto intro!: SAT_VIO.VSinceInf simp: i_LTP_tau i_ETP_tau + split: enat.splits)} + moreover + { assume unsat: "\ \\, v, i\ \ \ \<^bold>S I \" + and jw: "\j\i. mem (\ \ i j) I \ \\, v, j\ \ \" + and i_ge: "\ \ 0 + left I \ \ \ i" + from unsat jw not_sat_SinceD[of \ v i \ I \] + obtain j where j: "j \ i" + "case right I of \ \ True | enat n \ ETP \ (\ \ i - n) \ j" + "\ \\, v, j\ \ \" "(\k \ {j .. (min i (LTP \ (\ \ i - left I)))}. + \ \\, v, k\ \ \)" by (auto split: enat.splits) + with Since have "VIO \ v i (\ \<^bold>S I \)" + using i_ge unsat jw + by (auto intro!: SAT_VIO.VSince) } + ultimately show ?case + by (force simp del: sat.simps) +next + case (Until \ I \) + from \_mono have i0: "\ \ 0 \ \ \ i" by auto + { assume "\\, v, i\ \ \ \<^bold>U I \" + with Until have "SAT \ v i (\ \<^bold>U I \)" + by (auto intro: SAT_VIO.SUntil) } + moreover + { assume unsat: "\ \\, v, i\ \ \ \<^bold>U I \" + and witness: "\j \ i. mem (\ \ j i) I \ \\, v, j\ \ \" + from this Until not_sat_UntilD[of \ v i \ I \] obtain j + where j: "j \ i" "(case right I of \ \ True | enat n + \ j < LTP \ (\ \ i + n))" "\ (\\, v, j\ \ \)" + "(\k \ {(max i (ETP \ (\ \ i + left I))) .. j}. \ \\, v, k\ \ \)" + by auto + with Until have "VIO \ v i (\ \<^bold>U I \)" + using unsat witness + by (auto intro!: SAT_VIO.VUntil) } + moreover + { assume unsat: "\ \\, v, i\ \ \ \<^bold>U I \" + and no_witness: "\j \ i. \ mem (\ \ j i) I \ \ \\, v, j\ \ \" + with Until have "VIO \ v i (\ \<^bold>U I \)" + by (auto intro!: SAT_VIO.VUntilInf simp: add_increasing2 i0 i_LTP_tau i_ETP_tau + split: enat.splits) + } + ultimately show ?case by auto +qed (auto intro: SAT_VIO.intros) + +lemmas completeness = completeness_raw[THEN conjunct1, THEN mp] completeness_raw[THEN conjunct2, THEN mp] + +lemma SAT_or_VIO: "SAT \ v i \ \ VIO \ v i \" + using completeness[of \ v i \] by auto + +end + +unbundle MFOTL_no_notation + +(*<*) +end +(*>*) \ No newline at end of file diff --git a/thys/MFOTL_Checker/ROOT b/thys/MFOTL_Checker/ROOT new file mode 100644 --- /dev/null +++ b/thys/MFOTL_Checker/ROOT @@ -0,0 +1,21 @@ +chapter AFP + +session MFOTL_Checker (AFP) = "MFOTL_Monitor" + + options [timeout = 1200] + theories + Trace + Formula + Partition + Proof_System + Proof_Object + Prelim + Checker + Event_Data + Checker_Code + Monitor + Examples + document_files + "root.tex" + "root.bib" + export_files (in ".") [2] + "MFOTL_Checker.Checker_Code:code/**" diff --git a/thys/MFOTL_Checker/Trace.thy b/thys/MFOTL_Checker/Trace.thy new file mode 100644 --- /dev/null +++ b/thys/MFOTL_Checker/Trace.thy @@ -0,0 +1,449 @@ +(*<*) +theory Trace + imports "MFOTL_Monitor.Interval" "HOL-Library.Stream" +begin +(*>*) + +section \Traces and Trace Prefixes\ + +subsection \Infinite Traces\ + +coinductive ssorted :: "'a :: linorder stream \ bool" where + "shd s \ shd (stl s) \ ssorted (stl s) \ ssorted s" + +lemma ssorted_siterate[simp]: "(\n. n \ f n) \ ssorted (siterate f n)" + by (coinduction arbitrary: n) auto + +lemma ssortedD: "ssorted s \ s !! i \ stl s !! i" + by (induct i arbitrary: s) (auto elim: ssorted.cases) + +lemma ssorted_sdrop: "ssorted s \ ssorted (sdrop i s)" + by (coinduction arbitrary: i s) (auto elim: ssorted.cases ssortedD) + +lemma ssorted_monoD: "ssorted s \ i \ j \ s !! i \ s !! j" +proof (induct "j - i" arbitrary: j) + case (Suc x) + from Suc(1)[of "j - 1"] Suc(2-4) ssortedD[of s "j - 1"] + show ?case by (cases j) (auto simp: le_Suc_eq Suc_diff_le) +qed simp + +lemma sorted_stake: "ssorted s \ sorted (stake i s)" + by (induct i arbitrary: s) + (auto elim: ssorted.cases simp: in_set_conv_nth + intro!: ssorted_monoD[of _ 0, simplified, THEN order_trans, OF _ ssortedD]) + +lemma ssorted_monoI: "\i j. i \ j \ s !! i \ s !! j \ ssorted s" + by (coinduction arbitrary: s) + (auto dest: spec2[of _ "Suc _" "Suc _"] spec2[of _ 0 "Suc 0"]) + +lemma ssorted_iff_mono: "ssorted s \ (\i j. i \ j \ s !! i \ s !! j)" + using ssorted_monoI ssorted_monoD by metis + +lemma ssorted_iff_le_Suc: "ssorted s \ (\i. s !! i \ s !! Suc i)" + using mono_iff_le_Suc[of "snth s"] by (simp add: mono_def ssorted_iff_mono) + +definition "sincreasing s = (\x. \i. x < s !! i)" + +lemma sincreasingI: "(\x. \i. x < s !! i) \ sincreasing s" + by (simp add: sincreasing_def) + +lemma sincreasing_grD: + fixes x :: "'a :: semilattice_sup" + assumes "sincreasing s" + shows "\j>i. x < s !! j" +proof - + let ?A = "insert x {s !! n | n. n \ i}" + from assms obtain j where *: "Sup_fin ?A < s !! j" + by (auto simp: sincreasing_def) + then have "x < s !! j" + by (rule order.strict_trans1[rotated]) (auto intro: Sup_fin.coboundedI) + moreover have "i < j" + proof (rule ccontr) + assume "\ i < j" + then have "s !! j \ ?A" by (auto simp: not_less) + then have "s !! j \ Sup_fin ?A" + by (auto intro: Sup_fin.coboundedI) + with * show False by simp + qed + ultimately show ?thesis by blast +qed + +lemma sincreasing_siterate_nat[simp]: + fixes n :: nat + assumes "(\n. n < f n)" + shows "sincreasing (siterate f n)" +unfolding sincreasing_def proof + fix x + show "\i. x < siterate f n !! i" + proof (induction x) + case 0 + have "0 < siterate f n !! 1" + using order.strict_trans1[OF le0 assms] by simp + then show ?case .. + next + case (Suc x) + then obtain i where "x < siterate f n !! i" .. + then have "Suc x < siterate f n !! Suc i" + using order.strict_trans1[OF _ assms] by (simp del: snth.simps) + then show ?case .. + qed +qed + +lemma sincreasing_stl: "sincreasing s \ sincreasing (stl s)" for s :: "'a :: semilattice_sup stream" + by (auto 0 4 simp: gr0_conv_Suc intro!: sincreasingI dest: sincreasing_grD[of s 0]) + +definition "sfinite s = (\i. finite (s !! i))" + +lemma sfiniteI: "(\i. finite (s !! i)) \ sfinite s" + by (simp add: sfinite_def) + +typedef 'a trace = "{s :: ('a set \ nat) stream. ssorted (smap snd s) \ sincreasing (smap snd s) \ sfinite (smap fst s)}" + by (intro exI[of _ "smap (\i. ({}, i)) nats"]) + (auto simp: stream.map_comp stream.map_ident sfinite_def cong: stream.map_cong) + +setup_lifting type_definition_trace + +lift_definition \ :: "'a trace \ nat \ 'a set" is + "\s i. fst (s !! i)" . +lift_definition \ :: "'a trace \ nat \ nat" is + "\s i. snd (s !! i)" . + +lemma stream_eq_iff: "s = s' \ (\n. s !! n = s' !! n)" + by (metis stream.map_cong0 stream_smap_nats) + +lemma trace_eqI: "(\i. \ \ i = \ \' i) \ (\i. \ \ i = \ \' i) \ \ = \'" + by transfer (auto simp: stream_eq_iff intro!: prod_eqI) + +lemma \_mono[simp]: "i \ j \ \ s i \ \ s j" + by transfer (auto simp: ssorted_iff_mono) + +lemma ex_le_\: "\j\i. x \ \ s j" + by (transfer fixing: i x) (auto dest!: sincreasing_grD[of _ i x] less_imp_le) + +lemma le_\_less: "\ \ i \ \ \ j \ j < i \ \ \ i = \ \ j" + by (simp add: antisym) + +lemma less_\D: "\ \ i < \ \ j \ i < j" + by (meson \_mono less_le_not_le not_le_imp_less) + +abbreviation "\ s i \ \ s i - \ s (i - 1)" + +subsection \Finite Trace Prefixes\ + +typedef 'a prefix = "{p :: ('a set \ nat) list. sorted (map snd p)}" + by (auto intro!: exI[of _ "[]"]) + +setup_lifting type_definition_prefix + +lift_definition pmap_\ :: "('a set \ 'b set) \ 'a prefix \ 'b prefix" is + "\f. map (\(x, i). (f x, i))" + by (simp add: split_beta comp_def) + +lift_definition last_ts :: "'a prefix \ nat" is + "\p. (case p of [] \ 0 | _ \ snd (last p))" . + +lift_definition first_ts :: "nat \ 'a prefix \ nat" is + "\n p. (case p of [] \ n | _ \ snd (hd p))" . + +lift_definition pnil :: "'a prefix" is "[]" by simp + +lift_definition plen :: "'a prefix \ nat" is "length" . + +lift_definition psnoc :: "'a prefix \ 'a set \ nat \ 'a prefix" is + "\p x. if (case p of [] \ 0 | _ \ snd (last p)) \ snd x then p @ [x] else []" +proof (goal_cases sorted_psnoc) + case (sorted_psnoc p x) + then show ?case + by (induction p) (auto split: if_splits list.splits) +qed + +instantiation prefix :: (type) order begin + +lift_definition less_eq_prefix :: "'a prefix \ 'a prefix \ bool" is + "\p q. \r. q = p @ r" . + +definition less_prefix :: "'a prefix \ 'a prefix \ bool" where + "less_prefix x y = (x \ y \ \ y \ x)" + +instance +proof (standard, goal_cases less refl trans antisym) + case (less x y) + then show ?case unfolding less_prefix_def .. +next + case (refl x) + then show ?case by transfer auto +next + case (trans x y z) + then show ?case by transfer auto +next + case (antisym x y) + then show ?case by transfer auto +qed + +end + +lemma psnoc_inject[simp]: + "last_ts p \ snd x \ last_ts q \ snd y \ psnoc p x = psnoc q y \ (p = q \ x = y)" + by transfer auto + +lift_definition prefix_of :: "'a prefix \ 'a trace \ bool" is "\p s. stake (length p) s = p" . + +lemma prefix_of_pnil[simp]: "prefix_of pnil \" + by transfer auto + +lemma plen_pnil[simp]: "plen pnil = 0" + by transfer auto + +lemma plen_mono: "\ \ \' \ plen \ \ plen \'" + by transfer auto + +lemma prefix_of_psnocE: "prefix_of (psnoc p x) s \ last_ts p \ snd x \ + (prefix_of p s \ \ s (plen p) = fst x \ \ s (plen p) = snd x \ P) \ P" + by transfer (simp del: stake.simps add: stake_Suc) + +lemma le_pnil[simp]: "pnil \ \" + by transfer auto + +lift_definition take_prefix :: "nat \ 'a trace \ 'a prefix" is stake + by (auto dest: sorted_stake) + +lemma plen_take_prefix[simp]: "plen (take_prefix i \) = i" + by transfer auto + +lemma plen_psnoc[simp]: "last_ts \ \ snd x \ plen (psnoc \ x) = plen \ + 1" + by transfer auto + +lemma prefix_of_take_prefix[simp]: "prefix_of (take_prefix i \) \" + by transfer auto + +lift_definition pdrop :: "nat \ 'a prefix \ 'a prefix" is drop + by (auto simp: drop_map[symmetric] sorted_wrt_drop) + +lemma pdrop_0[simp]: "pdrop 0 \ = \" + by transfer auto + +lemma prefix_of_antimono: "\ \ \' \ prefix_of \' s \ prefix_of \ s" + by transfer (auto simp del: stake_add simp add: stake_add[symmetric]) + +lemma prefix_of_imp_linear: "prefix_of \ \ \ prefix_of \' \ \ \ \ \' \ \' \ \" +proof transfer + fix \ \' and \ :: "('a set \ nat) stream" + assume assms: "stake (length \) \ = \" "stake (length \') \ = \'" + show "(\r. \' = \ @ r) \ (\r. \ = \' @ r)" + proof (cases "length \" "length \'" rule: le_cases) + case le + then have "\' = take (length \) \' @ drop (length \) \'" + by simp + moreover have "take (length \) \' = \" + using assms le by (metis min.absorb1 take_stake) + ultimately show ?thesis by auto + next + case ge + then have "\ = take (length \') \ @ drop (length \') \" + by simp + moreover have "take (length \') \ = \'" + using assms ge by (metis min.absorb1 take_stake) + ultimately show ?thesis by auto + qed +qed + +lemma \_prefix_conv: "prefix_of p s \ prefix_of p s' \ i < plen p \ \ s i = \ s' i" + by transfer (simp add: stake_nth[symmetric]) + +lemma \_prefix_conv: "prefix_of p s \ prefix_of p s' \ i < plen p \ \ s i = \ s' i" + by transfer (simp add: stake_nth[symmetric]) + +lemma sincreasing_sdrop: + fixes s :: "('a :: semilattice_sup) stream" + assumes "sincreasing s" + shows "sincreasing (sdrop n s)" +proof (rule sincreasingI) + fix x + obtain i where "n < i" and "x < s !! i" + using sincreasing_grD[OF assms] by blast + then have "x < sdrop n s !! (i - n)" + by (simp add: sdrop_snth) + then show "\i. x < sdrop n s !! i" .. +qed + +lemma ssorted_shift: + "ssorted (xs @- s) = (sorted xs \ ssorted s \ (\x\set xs. \y\sset s. x \ y))" +proof safe + assume *: "ssorted (xs @- s)" + then show "sorted xs" + by (auto simp: ssorted_iff_mono shift_snth sorted_iff_nth_mono split: if_splits) + from ssorted_sdrop[OF *, of "length xs"] show "ssorted s" + by (auto simp: sdrop_shift) + fix x y assume "x \ set xs" "y \ sset s" + then obtain i j where "i < length xs" "xs ! i = x" "s !! j = y" + by (auto simp: set_conv_nth sset_range) + with ssorted_monoD[OF *, of i "j + length xs"] show "x \ y" by auto +next + assume "sorted xs" "ssorted s" "\x\set xs. \y\sset s. x \ y" + then show "ssorted (xs @- s)" + proof (coinduction arbitrary: xs s) + case (ssorted xs s) + with \ssorted s\ show ?case + by (subst (asm) ssorted.simps) (auto 0 4 simp: neq_Nil_conv shd_sset intro: exI[of _ "_ # _"]) + qed +qed + +lemma sincreasing_shift: + assumes "sincreasing s" + shows "sincreasing (xs @- s)" +proof (rule sincreasingI) + fix x + from assms obtain i where "x < s !! i" + unfolding sincreasing_def by blast + then have "x < (xs @- s) !! (length xs + i)" + by simp + then show "\i. x < (xs @- s) !! i" .. +qed + +lift_definition pts :: "'a prefix \ nat list" is "map snd" . + +lemma pts_pmap_\[simp]: "pts (pmap_\ f \) = pts \" + by (transfer fixing: f) (simp add: split_beta) + +subsection \Earliest and Latest Time-Points\ + +definition ETP:: "'a trace \ nat \ nat" where + "ETP \ t = (LEAST i. \ \ i \ t)" + +definition LTP:: "'a trace \ nat \ nat" where + "LTP \ t = Max {i. (\ \ i) \ t}" + +abbreviation "\ \ i j \ (\ \ i - \ \ j)" + +abbreviation "ETP_p \ i b \ ETP \ ((\ \ i) - b)" +abbreviation "LTP_p \ i I \ min i (LTP \ ((\ \ i) - left I))" +abbreviation "ETP_f \ i I \ max i (ETP \ ((\ \ i) + left I))" +abbreviation "LTP_f \ i b \ LTP \ ((\ \ i) + b)" + +definition max_opt where + "max_opt a b = (case (a,b) of (Some x, Some y) \ Some (max x y) | _ \ None)" + +definition "LTP_p_safe \ i I = (if \ \ i - left I \ \ \ 0 then LTP_p \ i I else 0)" + +lemma i_ETP_tau: "i \ ETP \ n \ \ \ i \ n" +proof + assume P: "i \ ETP \ n" + define j where j_def: "j \ ETP \ n" + then have i_j: "\ \ i \ \ \ j" using P by auto + from j_def have "\ \ j \ n" + unfolding ETP_def using LeastI_ex ex_le_\ by force + then show "\ \ i \ n" using i_j by auto +next + assume Q: "\ \ i \ n" + then show "ETP \ n \ i" unfolding ETP_def + by (auto simp add: Least_le) +qed + +lemma tau_LTP_k: + assumes "\ \ 0 \ n" "LTP \ n < k" + shows "\ \ k > n" +proof - + have "finite {i. \ \ i \ n}" + by (rule ccontr, unfold infinite_nat_iff_unbounded_le mem_Collect_eq) + (metis Suc_le_eq i_ETP_tau leD) + then show ?thesis + using assms(2) Max.coboundedI linorder_not_less + unfolding LTP_def by auto +qed + +lemma i_LTP_tau: + assumes n_asm: "n \ \ \ 0" + shows "(i \ LTP \ n \ \ \ i \ n)" +proof + define A and j where A_def: "A \ {i. \ \ i \ n}" and j_def: "j \ LTP \ n" + assume P: "i \ LTP \ n" + from n_asm A_def have A_ne: "A \ {}" by auto + from j_def have i_j: "\ \ i \ \ \ j" using P by auto + have not_in: "k \ A" if "j < k" for k + using n_asm that tau_LTP_k leD + unfolding A_def j_def by blast + then have "A \ {0.. \ j \ n" + using Max_in[of A] A_def fin_A + unfolding LTP_def + by simp + then show "\ \ i \ n" using i_j by auto +next + define A and j where A_def: "A \ {i. \ \ i \ n}" and j_def: "j \ LTP \ n" + assume Q: "\ \ i \ n" + have not_in: "k \ A" if "j < k" for k + using n_asm that tau_LTP_k leD + unfolding A_def j_def by blast + then have "A \ {0.. A" using Q A_def by auto + ultimately show "i \ LTP \ n" + using Max_ge[of A] A_def + unfolding LTP_def + by auto +qed + +lemma ETP_\: "i \ ETP \ (\ \ l + n) \ \ \ i l \ n" +proof - + assume P: "i \ ETP \ (\ \ l + n)" + then have "\ \ i \ \ \ l + n" by (auto simp add: i_ETP_tau) + then show ?thesis by auto +qed + +lemma ETP_ge: "ETP \ (\ \ l + n + 1) > l" +proof - + define j where j_def: "j \ \ \ l + n + 1" + then have etp_j: "\ \ (ETP \ j) \ j" unfolding ETP_def + using LeastI_ex ex_le_\ by force + then have "\ \ (ETP \ j) > \ \ l" using j_def by auto + then show ?thesis using j_def less_\D by blast +qed + +lemma i_le_LTPi: "i \ LTP \ (\ \ i)" + using \_mono i_LTP_tau[of \ "\ \ i" i] + by auto + +lemma i_le_LTPi_add: "i \ LTP \ (\ \ i + n)" + using i_le_LTPi + by (simp add: add_increasing2 i_LTP_tau) + +lemma i_le_LTPi_minus: + assumes "\ \ 0 + n \ \ \ i" "i > 0" "n > 0" + shows "LTP \ (\ \ i - n) < i" + unfolding LTP_def +proof (subst Max_less_iff; (intro ballI; elim CollectE)?) + show "finite {j. \ \ j \ \ \ i - n}" + unfolding finite_nat_set_iff_bounded_le + proof (intro exI[of _ i], safe) + fix j + assume "\ \ j \ \ \ i - n" + with assms(1,3) show "j \ i" + by (metis add_leD2 add_strict_increasing le_add_diff_inverse less_\D less_or_eq_imp_le) + qed +next + from assms(1) show "{j. \ \ j \ \ \ i - n} \ {}" + by (auto simp: le_diff_conv2) +next + fix j + assume "\ \ j \ \ \ i - n" + with assms(1,3) show "j < i" + by (metis add_leD2 add_strict_increasing le_add_diff_inverse less_\D) +qed + +lemma i_ge_etpi: "ETP \ (\ \ i) \ i" + using i_ETP_tau by auto + +(*<*) +end +(*>*) diff --git a/thys/MFOTL_Checker/document/root.bib b/thys/MFOTL_Checker/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/MFOTL_Checker/document/root.bib @@ -0,0 +1,34 @@ +@inproceedings{DBLP:conf/tacas/LimaHRTY23, + author = {Leonardo Lima and + Andrei Herasimau and + Martin Raszyk and + Dmitriy Traytel and + Simon Yuan}, + editor = {Sriram Sankaranarayanan and + Natasha Sharygina}, + title = {Explainable Online Monitoring of Metric Temporal Logic}, + booktitle = {{TACAS} 2023}, + series = {LNCS}, + volume = {13994}, + pages = {473--491}, + publisher = {Springer}, + year = {2023}, + doi = {10.1007/978-3-031-30820-8_28} +} + +@inproceedings{DBLP:conf/tacas/LimaHT24, + author = {Leonardo Lima and + Jonathan Juli{\'a}n Huerta y Munive and + Dmitriy Traytel}, + editor = {Bernd Finkbeiner and + Laura Kov{\'a}cs}, + title = {Explainable Online Monitoring of Metric First-Order Temporal Logic}, + booktitle = {{TACAS} 2024}, + series = {LNCS}, + volume = {14570}, + pages = {288--307}, + publisher = {Springer}, + year = {2024}, + doi = {10.1007/978-3-031-57246-3_16} +} + diff --git a/thys/MFOTL_Checker/document/root.tex b/thys/MFOTL_Checker/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/MFOTL_Checker/document/root.tex @@ -0,0 +1,44 @@ +\documentclass[10pt,a4paper]{article} +\usepackage[T1]{fontenc} +\usepackage{isabelle,isabellesym} +\usepackage{a4wide} +\usepackage[english]{babel} +\usepackage{eufrak} +\usepackage{amssymb} + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{literal} + + +\begin{document} + +\title{A Verified Proof Checker for Metric First-Order Temporal Logic} +\author{Andrei Herasimau \and Jonathan Jul\'{i}an Huerta y Munive \and Leonardo Lima \and Martin Raszyk \and Dmitriy Traytel} + +\maketitle + +\begin{abstract} +Metric first-order temporal logic (MFOTL) is an expressive formalism for specifying temporal and data-dependent constraints on streams of time-stamped, data-carrying events. Recently, we have developed a monitoring algorithm that not only outputs the satisfaction or violation of MFOTL formulas but also explains its verdicts in the form of proof trees~\cite{DBLP:conf/tacas/LimaHRTY23,DBLP:conf/tacas/LimaHT24}. These explanations serve as certificates, and in this entry we verify the correctness of a certificate checker. The checker is used to certify the output of our new, unverified monitoring tool WhyMon. The formalization contains another unverified, executable implementation of an explanation-producing monitoring algorithm used to exemplify our checker. +\end{abstract} + +\tableofcontents + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% generated text of all theories +\input{session} + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,813 +1,814 @@ ABY3_Protocols ADS_Functor AI_Planning_Languages_Semantics AODV AOT AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Ackermanns_not_PR Actuarial_Mathematics Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximate_Model_Counting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BD_Security_Compositional BNF_CC BNF_Operations BTree Balog_Szemeredi_Gowers Banach_Steinhaus Belief_Revision Bell_Numbers_Spivey BenOr_Kozen_Reif Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binary_Code_Imprimitive Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Birkhoff_Finite_Distributive_Lattices Blue_Eyes Bondy Boolean_Expression_Checkers Boolos_Curious_Inference Boolos_Curious_Inference_Automated Bounded_Deducibility_Security Broadcast_Psi Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CHERI-C_Memory_Model CISC-Kernel Concentration_Inequalities CRDT CRYSTALS-Kyber CRYSTALS-Kyber_Security CSP_RefTK CVP_Hardness CYK CZH_Elementary_Categories CZH_Foundations CZH_Universal_Constructions CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cardinality_Continuum Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Catoids Cauchy Cayley_Hamilton Certification_Monads Ceva Chandy_Lamport Chebyshev_Polynomials Chord_Segments Circus Clean Clique_and_Monotone_Circuits ClockSynchInst Closest_Pair_Points CoCon CoSMeDis CoSMed CofGroups Coinductive Coinductive_Languages Collections Combinable_Wands Combinatorial_Enumeration_Algorithms Combinatorics_Words Combinatorics_Words_Graph_Lemma Combinatorics_Words_Lyndon CommCSL Commuting_Hermitian Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Bounded_Operators Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentHOL ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Conditional_Simplification Conditional_Transfer_Rule CondNormReasHOL Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers Continued_Fractions Cook_Levin CoreC++ Core_DOM Core_SC_DOM Correctness_Algebras Cotangent_PFD_Formula Count_Complex_Roots Coupledsim_Contrasim CryptHOL CryptoBasedCompositionalProperties Crypto_Standards CubicalCategories Cubic_Quartic_Equations DCR-ExecutionEquivalence DFS_Framework DOM_Components DPRM_Theorem DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Dedekind_Real Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Design_Theory Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic DigitsInBase Digit_Expansions Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Directed_Sets Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation Disintegration DiskPaxos Distributed_Distinct_Elements Dominance_CHK Doob_Convergence DynamicArchitectures Dynamic_Tables E_Transcendental Earley_Parser Echelon_Form EdmondsKarp_Maxflow Edwards_Elliptic_Curves_Group Efficient-Mergesort Efficient_Weighted_Path_Order Elimination_Of_Repeated_Factors Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Equivalence_Relation_Enumeration Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Euler_Polyhedron_Formula Eudoxus_Reals Eval_FO Example-Submission Executable_Randomized_Algorithms Expander_Graphs Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Axiomatic FOL_Harrison FOL_Seq_Calc1 FOL_Seq_Calc2 FOL_Seq_Calc3 FO_Theory_Rewriting FSM_Tests Factor_Algebraic_Polynomial Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF Finite_Fields Finitely_Generated_Abelian_Groups First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Fishers_Inequality Fixed_Length_Vector Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Foundation_of_geometry Fourier Free-Boolean-Algebra Free-Groups Frequency_Moments Fresh_Identifiers FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC GaleStewart_Games Gale_Shapley Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic Given_Clause_Loops Go GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Gray_Codes Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Grothendieck_Schemes Group-Ring-Module HOL-CSP HOL-CSPM HOL-CSP_OpSem HOLCF-Prelude HRB-Slicing Hahn_Jordan_Decomposition Hales_Jewett Heard_Of Hello_World HereditarilyFinite Hermite Hermite_Lindemann Hidden_Markov_Models Higher_Order_Terms HoareForDivergence Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL HyperHoareLogic Hyperdual Hypergraph_Basics Hypergraph_Colourings IEEE_Floating_Point IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IMP_Compiler IMP_Compiler_Reuse Interval_Analysis IO_Language_Conformance IP_Addresses Imperative_Insertion_Sort Implicational_Logic Impossible_Geometry IMP_Noninterference Incompleteness Incredible_Proof_Machine Independence_CH Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpolation_Polynomials_HOL_Algebra Interpreter_Optimizations Interval_Arithmetic_Word32 Intro_Dest_Elim Involutions2Squares Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Irrationals_From_THEBOOK IsaGeoCoq IsaNet Isabelle_C Isabelle_hoops Isabelle_Marries_Dirac Isabelle_Meta_Model Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD Karatsuba KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Khovanskii_Theorem Kleene_Algebra Kneser_Cauchy_Davenport Knights_Tour Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt KnuthMorrisPratt Koenigsberg_Friendship Kruskal Kummer_Congruence Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LP_Duality LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_Series Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Laws_of_Large_Numbers Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option Lifting_the_Exponent LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Logging_Independent_Anonymity Lovasz_Local Lowe_Ontological_Argument Lower_Semicontinuous Lp Lucas_Theorem MDP-Algorithms MDP-Rewards MFMC_Countable MFODL_Monitor_Optimized +MFOTL_Checker MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Martingales Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Maximum_Segment_Sum Median_Method Median_Of_Medians_Selection Menger Mereology Mersenne_Primes Metalogic_ProofChecker MHComputation MiniML MiniSail Minimal_SSA Minkowskis_Theorem Minsky_Machines MLSS_Decision_Proc ML_Unification Modal_Logics_for_NTS Modular_Assembly_Kit_Security Modular_arithmetic_LLL_and_HNF_algorithms Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multi_Party_Computation Multirelations Multirelations_Heterogeneous Multiset_Ordering_NPC Multitape_To_Singletape_TM Myhill-Nerode Name_Carrying_Type_Inference Nano_JSON Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers No_FTL_observers_Gen_Rel Nominal2 Nominal_Myhill_Nerode Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Number_Theoretic_Transform Octonions OmegaCatoidsQuantales OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker PAL PAPP_Impossibility PCF PLM POPLmark-deBruijn PSemigroupsConvolution Package_logic Padic_Field Padic_Ints Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect_Fields Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Pluennecke_Ruzsa_Inequality Poincare_Bendixson Poincare_Disc Polygonal_Number_Theorem Polylog Polynomial_Crit_Geometry Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Prefix_Free_Code_Combinators Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Probability_Inequality_Completeness Program-Conflict-Analysis Progress_Tracking Projective_Geometry Projective_Measurements Promela Proof_Strategy_Language PropResPI Propositional_Logic_Class Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem Public_Announcement_Logic Q0_Metatheory Q0_Soundness QBF_Solver_Verification QHLProver QR_Decomposition Quantales Quantales_Converse Quantifier_Elimination_Hybrid Quasi_Borel_Spaces Quaternions Query_Optimization Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Real_Power Real_Time_Deque Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Region_Quadtrees Registers Regression_Test_Selection Regular-Sets Regular_Algebras Regular_Tree_Relations Relation_Algebra Relational-Incorrectness-Logic Relational_Cardinality Relational_Disjoint_Set_Forests Relational_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rensets Rep_Fin_Groups ResiduatedTransitionSystem Residuated_Lattices Resolution_FOL Rewrite_Properties_Reduction Rewriting_Z Ribbon_Proofs Risk_Free_Lending Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Roth_Arithmetic_Progressions Routing Roy_Floyd_Warshall SATSolverVerification SCC_Bloemen_Sequential SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 S_Finite_Measure_Monad Safe_Distance Safe_OCL Safe_Range_RC Saturation_Framework Saturation_Framework_Extensions Sauer_Shelah_Lemma Schutz_Spacetime Schwartz_Zippel Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL Separation_Logic_Unbounded SequentInvertibility Shadow_DOM Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Clause_Learning Simple_Firewall Simplex Simplicial_complexes_and_boolean_functions SimplifiedOntologicalArgument Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Solidity Sophomores_Dream Sort_Encodings Source_Coding_Theorem SpecCheck Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Stalnaker_Logic Standard_Borel_Spaces Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code StrictOmegaCategories Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras Sumcheck_Protocol SumSquares Sunflowers SuperCalc Suppes_Theorem Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Synthetic_Completeness Szemeredi_Regularity Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Three_Circles Three_Squares Timed_Automata Topological_Semantics Topology TortoiseHare TsirelsonBound Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Transitive_Models Transport Treaps Tree-Automata Tree_Decomposition Tree_Enumeration Triangle Trie Turans_Graph_Theorem Twelvefold_Way Two_Generated_Word_Monoids_Intersection Tycon Types_Tableaus_and_Goedels_God Types_To_Sets_Extension UPF UPF_Firewall UTP Uncertainty_Principle Undirected_Graph_Theory Universal_Hash_Families Universal_Turing_Machine UpDown_Scheme VYDRA_MDL Valuation Van_Emde_Boas_Trees Van_der_Waerden VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves Virtual_Substitution VolpanoSmith WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Weighted_Arithmetic_Geometric_Mean Weighted_Path_Order Well_Quasi_Orders Wetzels_Problem Wieferich_Kempner Winding_Number_Eval Word_Lib WorkerWrapper X86_Semantics XML Youngs_Inequality ZFC_in_HOL Zeckendorf Zeta_3_Irrational Zeta_Function pGCL Labeled_Transition_Systems Pushdown_Systems