Theory Window
theory Window
imports "HOL-Library.AList" "HOL-Library.Mapping" "HOL-Library.While_Combinator" Timestamp
begin
type_synonym ('a, 'b) mmap = "('a × 'b) list"
inductive chain_le :: "'d :: timestamp list ⇒ bool" where
chain_le_Nil: "chain_le []"
| chain_le_singleton: "chain_le [x]"
| chain_le_cons: "chain_le (y # xs) ⟹ x ≤ y ⟹ chain_le (x # y # xs)"
lemma chain_le_app: "chain_le (zs @ [z]) ⟹ z ≤ w ⟹ chain_le ((zs @ [z]) @ [w])"
apply (induction "zs @ [z]" arbitrary: zs rule: chain_le.induct)
apply (auto intro: chain_le.intros)[2]
subgoal for y xs x zs
apply (cases zs)
apply (auto)
apply (metis append.assoc append_Cons append_Nil chain_le_cons)
done
done
inductive reaches_on :: "('e ⇒ ('e × 'f) option) ⇒ 'e ⇒ 'f list ⇒ 'e ⇒ bool"
for run :: "'e ⇒ ('e × 'f) option" where
"reaches_on run s [] s"
| "run s = Some (s', v) ⟹ reaches_on run s' vs s'' ⟹ reaches_on run s (v # vs) s''"
lemma reaches_on_init_Some: "reaches_on r s xs s' ⟹ r s' ≠ None ⟹ r s ≠ None"
by (auto elim: reaches_on.cases)
lemma reaches_on_split: "reaches_on run s vs s' ⟹ i < length vs ⟹
∃s'' s'''. reaches_on run s (take i vs) s'' ∧ run s'' = Some (s''', vs ! i) ∧ reaches_on run s''' (drop (Suc i) vs) s'"
proof (induction s vs s' arbitrary: i rule: reaches_on.induct)
case (2 s s' v vs s'')
show ?case
using 2(1,2)
proof (cases i)
case (Suc n)
show ?thesis
using 2
by (fastforce simp: Suc intro: reaches_on.intros)
qed (auto intro: reaches_on.intros)
qed auto
lemma reaches_on_split': "reaches_on run s vs s' ⟹ i ≤ length vs ⟹
∃s'' . reaches_on run s (take i vs) s'' ∧ reaches_on run s'' (drop i vs) s'"
proof (induction s vs s' arbitrary: i rule: reaches_on.induct)
case (2 s s' v vs s'')
show ?case
using 2(1,2)
proof (cases i)
case (Suc n)
show ?thesis
using 2
by (fastforce simp: Suc intro: reaches_on.intros)
qed (auto intro: reaches_on.intros)
qed (auto intro: reaches_on.intros)
lemma reaches_on_split_app: "reaches_on run s (vs @ vs') s' ⟹
∃s''. reaches_on run s vs s'' ∧ reaches_on run s'' vs' s'"
using reaches_on_split'[where i="length vs", of run s "vs @ vs'" s']
by auto
lemma reaches_on_inj: "reaches_on run s vs t ⟹ reaches_on run s vs' t' ⟹
length vs = length vs' ⟹ vs = vs' ∧ t = t'"
apply (induction s vs t arbitrary: vs' t' rule: reaches_on.induct)
apply (auto elim: reaches_on.cases)[1]
subgoal for s s' v vs s'' vs' t'
apply (rule reaches_on.cases[of run s' vs s'']; rule reaches_on.cases[of run s vs' t'])
apply assumption+
apply auto[2]
apply fastforce
apply (metis length_0_conv list.discI)
apply (metis Pair_inject length_Cons nat.inject option.inject)
done
done
lemma reaches_on_split_last: "reaches_on run s (xs @ [x]) s'' ⟹
∃s'. reaches_on run s xs s' ∧ run s' = Some (s'', x)"
apply (induction s "xs @ [x]" s'' arbitrary: xs x rule: reaches_on.induct)
apply simp
subgoal for s s' v vs s'' xs x
by (cases vs rule: rev_cases) (fastforce elim: reaches_on.cases intro: reaches_on.intros)+
done
lemma reaches_on_rev_induct[consumes 1]: "reaches_on run s vs s' ⟹
(⋀s. P s [] s) ⟹
(⋀s s' v vs s''. reaches_on run s vs s' ⟹ P s vs s' ⟹ run s' = Some (s'', v) ⟹
P s (vs @ [v]) s'') ⟹
P s vs s'"
proof (induction vs arbitrary: s s' rule: rev_induct)
case (snoc x xs)
from snoc(2) obtain s'' where s''_def: "reaches_on run s xs s''" "run s'' = Some (s', x)"
using reaches_on_split_last
by fast
show ?case
using snoc(4)[OF s''_def(1) _ s''_def(2)] snoc(1)[OF s''_def(1) snoc(3,4)]
by auto
qed (auto elim: reaches_on.cases)
lemma reaches_on_app: "reaches_on run s vs s' ⟹ run s' = Some (s'', v) ⟹
reaches_on run s (vs @ [v]) s''"
by (induction s vs s' rule: reaches_on.induct) (auto intro: reaches_on.intros)
lemma reaches_on_trans: "reaches_on run s vs s' ⟹ reaches_on run s' vs' s'' ⟹
reaches_on run s (vs @ vs') s''"
by (induction s vs s' rule: reaches_on.induct) (auto intro: reaches_on.intros)
lemma reaches_onD: "reaches_on run s ((t, b) # vs) s' ⟹
∃s''. run s = Some (s'', (t, b)) ∧ reaches_on run s'' vs s'"
by (auto elim: reaches_on.cases)
lemma reaches_on_setD: "reaches_on run s vs s' ⟹ x ∈ set vs ⟹
∃vs' vs'' s''. reaches_on run s (vs' @ [x]) s'' ∧ reaches_on run s'' vs'' s' ∧ vs = vs' @ x # vs''"
proof (induction s vs s' rule: reaches_on_rev_induct)
case (2 s s' v vs s'')
show ?case
proof (cases "x ∈ set vs")
case True
obtain vs' vs'' s''' where split_def: "reaches_on run s (vs' @ [x]) s'''"
"reaches_on run s''' vs'' s'" "vs = vs' @ x # vs''"
using 2(3)[OF True]
by auto
show ?thesis
using split_def(1,3) reaches_on_app[OF split_def(2) 2(2)]
by auto
next
case False
have x_v: "x = v"
using 2(4) False
by auto
show ?thesis
unfolding x_v
using reaches_on_app[OF 2(1,2)] reaches_on.intros(1)[of run s'']
by auto
qed
qed auto
lemma reaches_on_len: "∃vs s'. reaches_on run s vs s' ∧ (length vs = n ∨ run s' = None)"
proof (induction n arbitrary: s)
case (Suc n)
show ?case
proof (cases "run s")
case (Some x)
obtain s' v where x_def: "x = (s', v)"
by (cases x) auto
obtain vs s'' where s''_def: "reaches_on run s' vs s''" "length vs = n ∨ run s'' = None"
using Suc[of s']
by auto
show ?thesis
using reaches_on.intros(2)[OF Some[unfolded x_def] s''_def(1)] s''_def(2)
by fastforce
qed (auto intro: reaches_on.intros)
qed (auto intro: reaches_on.intros)
lemma reaches_on_NilD: "reaches_on run q [] q' ⟹ q = q'"
by (auto elim: reaches_on.cases)
lemma reaches_on_ConsD: "reaches_on run q (x # xs) q' ⟹ ∃q''. run q = Some (q'', x) ∧ reaches_on run q'' xs q'"
by (auto elim: reaches_on.cases)
inductive reaches :: "('e ⇒ ('e × 'f) option) ⇒ 'e ⇒ nat ⇒ 'e ⇒ bool"
for run :: "'e ⇒ ('e × 'f) option" where
"reaches run s 0 s"
| "run s = Some (s', v) ⟹ reaches run s' n s'' ⟹ reaches run s (Suc n) s''"
lemma reaches_Suc_split_last: "reaches run s (Suc n) s' ⟹ ∃s'' x. reaches run s n s'' ∧ run s'' = Some (s', x)"
proof (induction n arbitrary: s)
case (Suc n)
obtain s'' x where s''_def: "run s = Some (s'', x)" "reaches run s'' (Suc n) s'"
using Suc(2)
by (auto elim: reaches.cases)
show ?case
using s''_def(1) Suc(1)[OF s''_def(2)]
by (auto intro: reaches.intros)
qed (auto elim!: reaches.cases intro: reaches.intros)
lemma reaches_invar: "reaches f x n y ⟹ P x ⟹ (⋀z z' v. P z ⟹ f z = Some (z', v) ⟹ P z') ⟹ P y"
by (induction x n y rule: reaches.induct) auto
lemma reaches_cong: "reaches f x n y ⟹ P x ⟹ (⋀z z' v. P z ⟹ f z = Some (z', v) ⟹ P z') ⟹ (⋀z. P z ⟹ f' (g z) = map_option (apfst g) (f z)) ⟹ reaches f' (g x) n (g y)"
by (induction x n y rule: reaches.induct) (auto intro: reaches.intros)
lemma reaches_on_n: "reaches_on run s vs s' ⟹ reaches run s (length vs) s'"
by (induction s vs s' rule: reaches_on.induct) (auto intro: reaches.intros)
lemma reaches_on: "reaches run s n s' ⟹ ∃vs. reaches_on run s vs s' ∧ length vs = n"
by (induction s n s' rule: reaches.induct) (auto intro: reaches_on.intros)
definition ts_at :: "('d × 'b) list ⇒ nat ⇒ 'd" where
"ts_at rho i = fst (rho ! i)"
definition bs_at :: "('d × 'b) list ⇒ nat ⇒ 'b" where
"bs_at rho i = snd (rho ! i)"
fun sub_bs :: "('d × 'b) list ⇒ nat × nat ⇒ 'b list" where
"sub_bs rho (i, j) = map (bs_at rho) [i..<j]"
definition steps :: "('c ⇒ 'b ⇒ 'c) ⇒ ('d × 'b) list ⇒ 'c ⇒ nat × nat ⇒ 'c" where
"steps step rho q ij = foldl step q (sub_bs rho ij)"
definition acc :: "('c ⇒ 'b ⇒ 'c) ⇒ ('c ⇒ bool) ⇒ ('d × 'b) list ⇒
'c ⇒ nat × nat ⇒ bool" where
"acc step accept rho q ij = accept (steps step rho q ij)"
definition sup_acc :: "('c ⇒ 'b ⇒ 'c) ⇒ ('c ⇒ bool) ⇒ ('d × 'b) list ⇒
'c ⇒ nat ⇒ nat ⇒ ('d × nat) option" where
"sup_acc step accept rho q i j =
(let L' = {l ∈ {i..<j}. acc step accept rho q (i, Suc l)}; m = Max L' in
if L' = {} then None else Some (ts_at rho m, m))"
definition sup_leadsto :: "'c ⇒ ('c ⇒ 'b ⇒ 'c) ⇒ ('d × 'b) list ⇒
nat ⇒ nat ⇒ 'c ⇒ 'd option" where
"sup_leadsto init step rho i j q =
(let L' = {l. l < i ∧ steps step rho init (l, j) = q}; m = Max L' in
if L' = {} then None else Some (ts_at rho m))"
definition mmap_keys :: "('a, 'b) mmap ⇒ 'a set" where
"mmap_keys kvs = set (map fst kvs)"
definition mmap_lookup :: "('a, 'b) mmap ⇒ 'a ⇒ 'b option" where
"mmap_lookup = map_of"
definition valid_s :: "'c ⇒ ('c ⇒ 'b ⇒ 'c) ⇒ ('c × 'b, 'c) mapping ⇒ ('c ⇒ bool) ⇒
('d × 'b) list ⇒ nat ⇒ nat ⇒ nat ⇒ ('c, 'c × ('d × nat) option) mmap ⇒ bool" where
"valid_s init step st accept rho u i j s ≡
(∀q bs. case Mapping.lookup st (q, bs) of None ⇒ True | Some v ⇒ step q bs = v) ∧
(mmap_keys s = {q. (∃l ≤ u. steps step rho init (l, i) = q)} ∧ distinct (map fst s) ∧
(∀q. case mmap_lookup s q of None ⇒ True
| Some (q', tstp) ⇒ steps step rho q (i, j) = q' ∧ tstp = sup_acc step accept rho q i j))"