Theory Monitor

(*<*)
theory Monitor
  imports
    Formula
    Optimized_Join
    "MFOTL_Monitor.Abstract_Monitor"
    "HOL-Library.While_Combinator"
    "HOL-Library.Mapping"
    "Deriving.Derive"
    "Generic_Join.Generic_Join_Correctness"
begin
(*>*)

section ‹Generic monitoring algorithm›

text ‹The algorithm defined here abstracts over the implementation of the temporal operators.›

subsection ‹Monitorable formulas›

definition "mmonitorable φ  safe_formula φ  Formula.future_bounded φ"
definition "mmonitorable_regex b g r  safe_regex b g r  Regex.pred_regex Formula.future_bounded r"

definition is_simple_eq :: "Formula.trm  Formula.trm  bool" where
  "is_simple_eq t1 t2 = (Formula.is_Const t1  (Formula.is_Const t2  Formula.is_Var t2) 
    Formula.is_Var t1  Formula.is_Const t2)"

fun mmonitorable_exec :: "Formula.formula  bool" where
  "mmonitorable_exec (Formula.Eq t1 t2) = is_simple_eq t1 t2"
| "mmonitorable_exec (Formula.Neg (Formula.Eq (Formula.Var x) (Formula.Var y))) = (x = y)"
| "mmonitorable_exec (Formula.Pred e ts) = list_all (λt. Formula.is_Var t  Formula.is_Const t) ts"
| "mmonitorable_exec (Formula.Let p φ ψ) = ({0..<Formula.nfv φ}  Formula.fv φ  mmonitorable_exec φ  mmonitorable_exec ψ)"
| "mmonitorable_exec (Formula.Neg φ) = (fv φ = {}  mmonitorable_exec φ)"
| "mmonitorable_exec (Formula.Or φ ψ) = (fv φ = fv ψ  mmonitorable_exec φ  mmonitorable_exec ψ)"
| "mmonitorable_exec (Formula.And φ ψ) = (mmonitorable_exec φ 
    (safe_assignment (fv φ) ψ  mmonitorable_exec ψ 
      fv ψ  fv φ  (is_constraint ψ  (case ψ of Formula.Neg ψ'  mmonitorable_exec ψ' | _  False))))"
| "mmonitorable_exec (Formula.Ands l) = (let (pos, neg) = partition mmonitorable_exec l in
    pos  []  list_all mmonitorable_exec (map remove_neg neg) 
    (set (map fv neg))  (set (map fv pos)))"
| "mmonitorable_exec (Formula.Exists φ) = (mmonitorable_exec φ)"
| "mmonitorable_exec (Formula.Agg y ω b f φ) = (mmonitorable_exec φ 
    y + b  Formula.fv φ  {0..<b}  Formula.fv φ  Formula.fv_trm f  Formula.fv φ)"
| "mmonitorable_exec (Formula.Prev I φ) = (mmonitorable_exec φ)"
| "mmonitorable_exec (Formula.Next I φ) = (mmonitorable_exec φ)"
| "mmonitorable_exec (Formula.Since φ I ψ) = (Formula.fv φ  Formula.fv ψ 
    (mmonitorable_exec φ  (case φ of Formula.Neg φ'  mmonitorable_exec φ' | _  False))  mmonitorable_exec ψ)"
| "mmonitorable_exec (Formula.Until φ I ψ) = (Formula.fv φ  Formula.fv ψ  right I   
    (mmonitorable_exec φ  (case φ of Formula.Neg φ'  mmonitorable_exec φ' | _  False))  mmonitorable_exec ψ)"
| "mmonitorable_exec (Formula.MatchP I r) = Regex.safe_regex Formula.fv (λg φ. mmonitorable_exec φ  (g = Lax  (case φ of Formula.Neg φ'  mmonitorable_exec φ' | _  False))) Past Strict r"
| "mmonitorable_exec (Formula.MatchF I r) = (Regex.safe_regex Formula.fv (λg φ. mmonitorable_exec φ  (g = Lax  (case φ of Formula.Neg φ'  mmonitorable_exec φ' | _  False))) Futu Strict r  right I  )"
| "mmonitorable_exec _ = False"

lemma cases_Neg_iff:
  "(case φ of formula.Neg ψ  P ψ | _  False)  (ψ. φ = formula.Neg ψ  P ψ)"
  by (cases φ) auto

lemma safe_formula_mmonitorable_exec: "safe_formula φ  Formula.future_bounded φ  mmonitorable_exec φ"
proof (induct φ rule: safe_formula.induct)
  case (8 φ ψ)
  then show ?case
    unfolding safe_formula.simps future_bounded.simps mmonitorable_exec.simps
    by (auto simp: cases_Neg_iff)
next
  case (9 φ ψ)
  then show ?case
    unfolding safe_formula.simps future_bounded.simps mmonitorable_exec.simps
    by (auto simp: cases_Neg_iff)
next
  case (10 l)
  from "10.prems"(2) have bounded: "Formula.future_bounded φ" if "φ  set l" for φ
    using that by (auto simp: list.pred_set)
  obtain poss negs where posnegs: "(poss, negs) = partition safe_formula l" by simp
  obtain posm negm where posnegm: "(posm, negm) = partition mmonitorable_exec l" by simp
  have "set poss  set posm"
  proof (rule subsetI)
    fix x assume "x  set poss"
    then have "x  set l" "safe_formula x" using posnegs by simp_all
    then have "mmonitorable_exec x" using "10.hyps"(1) bounded by blast
    then show "x  set posm" using x  set poss posnegm posnegs by simp
  qed
  then have "set negm  set negs" using posnegm posnegs by auto
  obtain "poss  []" "list_all safe_formula (map remove_neg negs)"
    "(xset negs. fv x)  (xset poss. fv x)"
    using "10.prems"(1) posnegs by simp
  then have "posm  []" using set poss  set posm by auto
  moreover have "list_all mmonitorable_exec (map remove_neg negm)"
  proof -
    let ?l = "map remove_neg negm"
    have "x. x  set ?l  mmonitorable_exec x"
    proof -
      fix x assume "x  set ?l"
      then obtain y where "y  set negm" "x = remove_neg y" by auto
      then have "y  set negs" using set negm  set negs by blast
      then have "safe_formula x"
        unfolding x = remove_neg y using list_all safe_formula (map remove_neg negs)
        by (simp add: list_all_def)
      show "mmonitorable_exec x"
      proof (cases "z. y = Formula.Neg z")
        case True
        then obtain z where "y = Formula.Neg z" by blast
        then show ?thesis
          using "10.hyps"(2)[OF posnegs refl] x = remove_neg y y  set negs posnegs bounded
            safe_formula x by fastforce
      next
        case False
        then have "remove_neg y = y" by (cases y) simp_all
        then have "y = x" unfolding x = remove_neg y by simp
        show ?thesis
          using "10.hyps"(1) y  set negs posnegs safe_formula x unfolding y = x
          by auto
      qed
    qed
    then show ?thesis by (simp add: list_all_iff)
  qed
  moreover have "(xset negm. fv x)  (xset posm. fv x)"
    using  (fv ` set negs)   (fv ` set poss) set poss  set posm set negm  set negs
    by fastforce
  ultimately show ?case using posnegm by simp
next
  case (15 φ I ψ)
  then show ?case
    unfolding safe_formula.simps future_bounded.simps mmonitorable_exec.simps
    by (auto simp: cases_Neg_iff)
next
  case (16 φ I ψ)
  then show ?case
    unfolding safe_formula.simps future_bounded.simps mmonitorable_exec.simps
    by (auto simp: cases_Neg_iff)
next
  case (17 I r)
  then show ?case
    by (auto elim!: safe_regex_mono[rotated] simp: cases_Neg_iff regex.pred_set)
next
  case (18 I r)
  then show ?case
    by (auto elim!: safe_regex_mono[rotated] simp: cases_Neg_iff regex.pred_set)
qed (auto simp add: is_simple_eq_def list.pred_set)

lemma safe_assignment_future_bounded: "safe_assignment X φ  Formula.future_bounded φ"
  unfolding safe_assignment_def by (auto split: formula.splits)

lemma is_constraint_future_bounded: "is_constraint φ  Formula.future_bounded φ"
  by (induction rule: is_constraint.induct) simp_all

lemma mmonitorable_exec_mmonitorable: "mmonitorable_exec φ  mmonitorable φ"
proof (induct φ rule: mmonitorable_exec.induct)
  case (7 φ ψ)
  then show ?case
    unfolding mmonitorable_def mmonitorable_exec.simps safe_formula.simps
    by (auto simp: cases_Neg_iff intro: safe_assignment_future_bounded is_constraint_future_bounded)
next
  case (8 l)
  obtain poss negs where posnegs: "(poss, negs) = partition safe_formula l" by simp
  obtain posm negm where posnegm: "(posm, negm) = partition mmonitorable_exec l" by simp
  have pos_monitorable: "list_all mmonitorable_exec posm" using posnegm by (simp add: list_all_iff)
  have neg_monitorable: "list_all mmonitorable_exec (map remove_neg negm)"
    using "8.prems" posnegm by (simp add: list_all_iff)
  have "set posm  set poss"
    using "8.hyps"(1) posnegs posnegm unfolding mmonitorable_def by auto
  then have "set negs  set negm"
    using posnegs posnegm by auto

  have "poss  []" using "8.prems" posnegm set posm  set poss by auto
  moreover have "list_all safe_formula (map remove_neg negs)"
    using neg_monitorable "8.hyps"(2)[OF posnegm refl] set negs  set negm
    unfolding mmonitorable_def by (auto simp: list_all_iff)
  moreover have " (fv ` set negm)   (fv ` set posm)"
    using "8.prems" posnegm by simp
  then have " (fv ` set negs)   (fv ` set poss)"
    using set posm  set poss set negs  set negm by fastforce
  ultimately have "safe_formula (Formula.Ands l)" using posnegs by simp
  moreover have "Formula.future_bounded (Formula.Ands l)"
  proof -
    have "list_all Formula.future_bounded posm"
      using pos_monitorable "8.hyps"(1) posnegm unfolding mmonitorable_def
      by (auto elim!: list.pred_mono_strong)
    moreover have fnegm: "list_all Formula.future_bounded (map remove_neg negm)"
      using neg_monitorable "8.hyps"(2) posnegm unfolding mmonitorable_def
      by (auto elim!: list.pred_mono_strong)
    then have "list_all Formula.future_bounded negm"
    proof -
      have "x. x  set negm  Formula.future_bounded x"
      proof -
        fix x assume "x  set negm"
        have "Formula.future_bounded (remove_neg x)" using fnegm x  set negm by (simp add: list_all_iff)
        then show "Formula.future_bounded x" by (cases x) auto
      qed
      then show ?thesis by (simp add: list_all_iff)
    qed
    ultimately have "list_all Formula.future_bounded l" using posnegm by (auto simp: list_all_iff)
    then show ?thesis by (auto simp: list_all_iff)
  qed
  ultimately show ?case unfolding mmonitorable_def ..
next
  case (13 φ I ψ)
  then show ?case
    unfolding mmonitorable_def mmonitorable_exec.simps safe_formula.simps
    by (fastforce simp: cases_Neg_iff)
next
  case (14 φ I ψ)
  then show ?case
    unfolding mmonitorable_def mmonitorable_exec.simps safe_formula.simps
    by (fastforce simp: one_enat_def cases_Neg_iff)
next
  case (15 I r)
  then show ?case
    by (auto elim!: safe_regex_mono[rotated] dest: safe_regex_safe[rotated]
        simp: mmonitorable_regex_def mmonitorable_def cases_Neg_iff regex.pred_set)
next
  case (16 I r)
  then show ?case
    by (auto 0 3 elim!: safe_regex_mono[rotated] dest: safe_regex_safe[rotated]
        simp: mmonitorable_regex_def mmonitorable_def cases_Neg_iff regex.pred_set)
qed (auto simp add: mmonitorable_def mmonitorable_regex_def is_simple_eq_def one_enat_def list.pred_set)

lemma monitorable_formula_code[code]: "mmonitorable φ = mmonitorable_exec φ"
  using mmonitorable_exec_mmonitorable safe_formula_mmonitorable_exec mmonitorable_def
  by blast

subsection ‹Handling regular expressions›

datatype mregex =
  MSkip nat
  | MTestPos nat
  | MTestNeg nat
  | MPlus mregex mregex
  | MTimes mregex mregex
  | MStar mregex

primrec ok where
  "ok _ (MSkip n) = True"
| "ok m (MTestPos n) = (n < m)"
| "ok m (MTestNeg n) = (n < m)"
| "ok m (MPlus r s) = (ok m r  ok m s)"
| "ok m (MTimes r s) = (ok m r  ok m s)"
| "ok m (MStar r) = ok m r"

primrec from_mregex where
  "from_mregex (MSkip n) _ = Regex.Skip n"
| "from_mregex (MTestPos n) φs = Regex.Test (φs ! n)"
| "from_mregex (MTestNeg n) φs = (if safe_formula (Formula.Neg (φs ! n))
    then Regex.Test (Formula.Neg (Formula.Neg (Formula.Neg (φs ! n))))
    else Regex.Test (Formula.Neg (φs ! n)))"
| "from_mregex (MPlus r s) φs = Regex.Plus (from_mregex r φs) (from_mregex s φs)"
| "from_mregex (MTimes r s) φs = Regex.Times (from_mregex r φs) (from_mregex s φs)"
| "from_mregex (MStar r) φs = Regex.Star (from_mregex r φs)"

primrec to_mregex_exec where
  "to_mregex_exec (Regex.Skip n) xs = (MSkip n, xs)"
| "to_mregex_exec (Regex.Test φ) xs = (if safe_formula φ then (MTestPos (length xs), xs @ [φ])
     else case φ of Formula.Neg φ'  (MTestNeg (length xs), xs @ [φ']) | _  (MSkip 0, xs))"
| "to_mregex_exec (Regex.Plus r s) xs =
     (let (mr, ys) = to_mregex_exec r xs; (ms, zs) = to_mregex_exec s ys
     in (MPlus mr ms, zs))"
| "to_mregex_exec (Regex.Times r s) xs =
     (let (mr, ys) = to_mregex_exec r xs; (ms, zs) = to_mregex_exec s ys
     in (MTimes mr ms, zs))"
| "to_mregex_exec (Regex.Star r) xs =
     (let (mr, ys) = to_mregex_exec r xs in (MStar mr, ys))"

primrec shift where
  "shift (MSkip n) k = MSkip n"
| "shift (MTestPos i) k = MTestPos (i + k)"
| "shift (MTestNeg i) k = MTestNeg (i + k)"
| "shift (MPlus r s) k = MPlus (shift r k) (shift s k)"
| "shift (MTimes r s) k = MTimes (shift r k) (shift s k)"
| "shift (MStar r) k = MStar (shift r k)"

primrec to_mregex where
  "to_mregex (Regex.Skip n) = (MSkip n, [])"
| "to_mregex (Regex.Test φ) = (if safe_formula φ then (MTestPos 0, [φ])
     else case φ of Formula.Neg φ'  (MTestNeg 0, [φ']) | _  (MSkip 0, []))"
| "to_mregex (Regex.Plus r s) =
     (let (mr, ys) = to_mregex r; (ms, zs) = to_mregex s
     in (MPlus mr (shift ms (length ys)), ys @ zs))"
| "to_mregex (Regex.Times r s) =
     (let (mr, ys) = to_mregex r; (ms, zs) = to_mregex s
     in (MTimes mr (shift ms (length ys)), ys @ zs))"
| "to_mregex (Regex.Star r) =
     (let (mr, ys) = to_mregex r in (MStar mr, ys))"

lemma shift_0: "shift r 0 = r"
  by (induct r) auto

lemma shift_shift: "shift (shift r k) j = shift r (k + j)"
  by (induct r) auto

lemma to_mregex_to_mregex_exec:
  "case to_mregex r of (mr, φs)  to_mregex_exec r xs = (shift mr (length xs), xs @ φs)"
  by (induct r arbitrary: xs)
    (auto simp: shift_shift ac_simps split: formula.splits prod.splits)

lemma to_mregex_to_mregex_exec_Nil[code]: "to_mregex r = to_mregex_exec r []"
  using to_mregex_to_mregex_exec[where xs="[]" and r = r] by (auto simp: shift_0)

lemma ok_mono: "ok m mr  m  n  ok n mr"
  by (induct mr) auto

lemma from_mregex_cong: "ok m mr  (i < m. xs ! i = ys ! i)  from_mregex mr xs = from_mregex mr ys"
  by (induct mr) auto

lemma not_Neg_cases:
  "(ψ. φ  Formula.Neg ψ)  (case φ of formula.Neg ψ  f ψ | _  x) = x"
  by (cases φ) auto

lemma to_mregex_exec_ok:
  "to_mregex_exec r xs = (mr, ys)  zs. ys = xs @ zs  set zs = atms r  ok (length ys) mr"
proof (induct r arbitrary: xs mr ys)
  case (Skip x)
  then show ?case by (auto split: if_splits prod.splits simp: atms_def elim: ok_mono)
next
  case (Test x)
  show ?case proof (cases "x'. x = Formula.Neg x'")
    case True
    with Test show ?thesis by (auto split: if_splits prod.splits simp: atms_def elim: ok_mono)
  next
    case False
    with Test show ?thesis by (auto split: if_splits prod.splits simp: atms_def not_Neg_cases elim: ok_mono)
  qed
next
  case (Plus r1 r2)
  then show ?case by (fastforce split: if_splits prod.splits formula.splits simp: atms_def elim: ok_mono)
next
  case (Times r1 r2)
  then show ?case by (fastforce split: if_splits prod.splits formula.splits simp: atms_def elim: ok_mono)
next
  case (Star r)
  then show ?case by (auto split: if_splits prod.splits simp: atms_def elim: ok_mono)
qed

lemma ok_shift: "ok (i + m) (Monitor.shift r i)  ok m r"
  by (induct r) auto

lemma to_mregex_ok: "to_mregex r = (mr, ys)  set ys = atms r  ok (length ys) mr"
proof (induct r arbitrary: mr ys)
  case (Skip x)
  then show ?case by (auto simp: atms_def elim: ok_mono split: if_splits prod.splits)
next
  case (Test x)
  show ?case proof (cases "x'. x = Formula.Neg x'")
    case True
    with Test show ?thesis by (auto split: if_splits prod.splits simp: atms_def elim: ok_mono)
  next
    case False
    with Test show ?thesis by (auto split: if_splits prod.splits simp: atms_def not_Neg_cases elim: ok_mono)
  qed
next
  case (Plus r1 r2)
  then show ?case by (fastforce simp: ok_shift atms_def elim: ok_mono split: if_splits prod.splits)
next
  case (Times r1 r2)
  then show ?case by (fastforce simp: ok_shift atms_def elim: ok_mono split: if_splits prod.splits)
next
  case (Star r)
  then show ?case by (auto simp: atms_def elim: ok_mono split: if_splits prod.splits)
qed

lemma from_mregex_shift: "from_mregex (shift r (length xs)) (xs @ ys) = from_mregex r ys"
  by (induct r) (auto simp: nth_append)

lemma from_mregex_to_mregex: "safe_regex m g r  case_prod from_mregex (to_mregex r) = r"
  by (induct r rule: safe_regex.induct)
    (auto dest: to_mregex_ok intro!: from_mregex_cong simp: nth_append from_mregex_shift cases_Neg_iff
      split: prod.splits modality.splits)

lemma from_mregex_eq: "safe_regex m g r  to_mregex r = (mr, φs)  from_mregex mr φs = r"
  using from_mregex_to_mregex[of m g r] by auto

lemma from_mregex_to_mregex_exec: "safe_regex m g r  case_prod from_mregex (to_mregex_exec r xs) = r"
proof (induct r arbitrary: xs rule: safe_regex_induct)
  case (Plus b g r s)
  from Plus(3) Plus(1)[of xs] Plus(2)[of "snd (to_mregex_exec r xs)"] show ?case
    by (auto split: prod.splits simp: nth_append dest: to_mregex_exec_ok
        intro!: from_mregex_cong[where m = "length (snd (to_mregex_exec r xs))"])
next
  case (TimesF g r s)
  from TimesF(3) TimesF(2)[of xs] TimesF(1)[of "snd (to_mregex_exec r xs)"] show ?case
    by (auto split: prod.splits simp: nth_append dest: to_mregex_exec_ok
        intro!: from_mregex_cong[where m = "length (snd (to_mregex_exec r xs))"])
next
  case (TimesP g r s)
  from TimesP(3) TimesP(1)[of xs] TimesP(2)[of "snd (to_mregex_exec r xs)"] show ?case
    by (auto split: prod.splits simp: nth_append dest: to_mregex_exec_ok
        intro!: from_mregex_cong[where m = "length (snd (to_mregex_exec r xs))"])
next
  case (Star b g r)
  from Star(2) Star(1)[of xs] show ?case
    by (auto split: prod.splits)
qed (auto split: prod.splits simp: cases_Neg_iff)


derive linorder mregex

subsubsection ‹LPD›

definition saturate where
  "saturate f = while (λS. f S  S) f"

lemma saturate_code[code]:
  "saturate f S = (let S' = f S in if S' = S then S else saturate f S')"
  unfolding saturate_def Let_def
  by (subst while_unfold) auto

definition "MTimesL r S = MTimes r ` S"
definition "MTimesR R s = (λr. MTimes r s) ` R"

primrec LPD where
  "LPD (MSkip n) = (case n of 0  {} | Suc m  {MSkip m})"
| "LPD (MTestPos φ) = {}"
| "LPD (MTestNeg φ) = {}"
| "LPD (MPlus r s) = (LPD r  LPD s)"
| "LPD (MTimes r s) = MTimesR (LPD r) s  LPD s"
| "LPD (MStar r) = MTimesR (LPD r) (MStar r)"

primrec LPDi where
  "LPDi 0 r = {r}"
| "LPDi (Suc i) r = (s  LPD r. LPDi i s)"

lemma LPDi_Suc_alt: "LPDi (Suc i) r = (s  LPDi i r. LPD s)"
  by (induct i arbitrary: r) fastforce+

definition "LPDs r = (i. LPDi i r)"

lemma LPDs_refl: "r  LPDs r"
  by (auto simp: LPDs_def intro: exI[of _ 0])
lemma LPDs_trans: "r  LPD s  s  LPDs t  r  LPDs t"
  by (force simp: LPDs_def LPDi_Suc_alt simp del: LPDi.simps intro: exI[of _ "Suc _"])

lemma LPDi_Test:
  "LPDi i (MSkip 0)  {MSkip 0}"
  "LPDi i (MTestPos φ)  {MTestPos φ}"
  "LPDi i (MTestNeg φ)  {MTestNeg φ}"
  by (induct i) auto

lemma LPDs_Test:
  "LPDs (MSkip 0)  {MSkip 0}"
  "LPDs (MTestPos φ)  {MTestPos φ}"
  "LPDs (MTestNeg φ)  {MTestNeg φ}"
  unfolding LPDs_def using LPDi_Test by blast+

lemma LPDi_MSkip: "LPDi i (MSkip n)  MSkip ` {i. i  n}"
  by (induct i arbitrary: n) (auto dest: set_mp[OF LPDi_Test(1)] simp: le_Suc_eq split: nat.splits)

lemma LPDs_MSkip: "LPDs (MSkip n)  MSkip ` {i. i  n}"
  unfolding LPDs_def using LPDi_MSkip by auto

lemma LPDi_Plus: "LPDi i (MPlus r s)  {MPlus r s}  LPDi i r  LPDi i s"
  by (induct i arbitrary: r s) auto

lemma LPDs_Plus: "LPDs (MPlus r s)  {MPlus r s}  LPDs r  LPDs s"
  unfolding LPDs_def using LPDi_Plus[of _ r s] by auto

lemma LPDi_Times:
  "LPDi i (MTimes r s)  {MTimes r s}  MTimesR (j  i. LPDi j r) s  (j  i. LPDi j s)"
proof (induct i arbitrary: r s)
  case (Suc i)
  show ?case
    by (fastforce simp: MTimesR_def dest: bspec[of _ _ "Suc _"] dest!: set_mp[OF Suc])
qed simp

lemma LPDs_Times: "LPDs (MTimes r s)  {MTimes r s}  MTimesR (LPDs r) s  LPDs s"
  unfolding LPDs_def using LPDi_Times[of _ r s] by (force simp: MTimesR_def)

lemma LPDi_Star: "j  i  LPDi j (MStar r)  {MStar r}  MTimesR (j  i. LPDi j r) (MStar r)"
proof (induct i arbitrary: j r)
  case (Suc i)
  from Suc(2) show ?case
    by (auto 0 5 simp: MTimesR_def image_iff le_Suc_eq
        dest: bspec[of _ _ "Suc 0"] bspec[of _ _ "Suc _"] set_mp[OF Suc(1)] dest!: set_mp[OF LPDi_Times])
qed simp

lemma LPDs_Star: "LPDs (MStar r)  {MStar r}  MTimesR (LPDs r) (MStar r)"
  unfolding LPDs_def using LPDi_Star[OF order_refl, of _ r] by (force simp: MTimesR_def)

lemma finite_LPDs: "finite (LPDs r)"
proof (induct r)
  case (MSkip n)
  then show ?case by (intro finite_subset[OF LPDs_MSkip]) simp
next
  case (MTestPos φ)
  then show ?case by (intro finite_subset[OF LPDs_Test(2)]) simp
next
  case (MTestNeg φ)
  then show ?case by (intro finite_subset[OF LPDs_Test(3)]) simp
next
  case (MPlus r s)
  then show ?case by (intro finite_subset[OF LPDs_Plus]) simp
next
  case (MTimes r s)
  then show ?case by (intro finite_subset[OF LPDs_Times]) (simp add: MTimesR_def)
next
  case (MStar r)
  then show ?case by (intro finite_subset[OF LPDs_Star]) (simp add: MTimesR_def)
qed

context begin

private abbreviation (input) "addLPD r  λS. insert r S  Set.bind (insert r S) LPD"

private lemma mono_addLPD: "mono (addLPD r)"
  unfolding mono_def Set.bind_def by auto

private lemma LPDs_aux1: "lfp (addLPD r)  LPDs r"
  by (rule lfp_induct[OF mono_addLPD], auto intro: LPDs_refl LPDs_trans simp: Set.bind_def)

private lemma LPDs_aux2: "LPDi i r  lfp (addLPD r)"
proof (induct i)
  case 0
  then show ?case
    by (subst lfp_unfold[OF mono_addLPD]) auto
next
  case (Suc i)
  then show ?case
    by (subst lfp_unfold[OF mono_addLPD]) (auto simp: LPDi_Suc_alt simp del: LPDi.simps)
qed

lemma LPDs_alt: "LPDs r = lfp (addLPD r)"
  using LPDs_aux1 LPDs_aux2 by (force simp: LPDs_def)

lemma LPDs_code[code]:
  "LPDs r = saturate (addLPD r) {}"
  unfolding LPDs_alt saturate_def
  by (rule lfp_while[OF mono_addLPD _ finite_LPDs, of r]) (auto simp: LPDs_refl LPDs_trans Set.bind_def)

end

subsubsection ‹RPD›

primrec RPD where
  "RPD (MSkip n) = (case n of 0  {} | Suc m  {MSkip m})"
| "RPD (MTestPos φ) = {}"
| "RPD (MTestNeg φ) = {}"
| "RPD (MPlus r s) = (RPD r  RPD s)"
| "RPD (MTimes r s) = MTimesL r (RPD s)  RPD r"
| "RPD (MStar r) = MTimesL (MStar r) (RPD r)"

primrec RPDi where
  "RPDi 0 r = {r}"
| "RPDi (Suc i) r = (s  RPD r. RPDi i s)"

lemma RPDi_Suc_alt: "RPDi (Suc i) r = (s  RPDi i r. RPD s)"
  by (induct i arbitrary: r) fastforce+

definition "RPDs r = (i. RPDi i r)"

lemma RPDs_refl: "r  RPDs r"
  by (auto simp: RPDs_def intro: exI[of _ 0])
lemma RPDs_trans: "r  RPD s  s  RPDs t  r  RPDs t"
  by (force simp: RPDs_def RPDi_Suc_alt simp del: RPDi.simps intro: exI[of _ "Suc _"])

lemma RPDi_Test:
  "RPDi i (MSkip 0)  {MSkip 0}"
  "RPDi i (MTestPos φ)  {MTestPos φ}"
  "RPDi i (MTestNeg φ)  {MTestNeg φ}"
  by (induct i) auto

lemma RPDs_Test:
  "RPDs (MSkip 0)  {MSkip 0}"
  "RPDs (MTestPos φ)  {MTestPos φ}"
  "RPDs (MTestNeg φ)  {MTestNeg φ}"
  unfolding RPDs_def using RPDi_Test by blast+

lemma RPDi_MSkip: "RPDi i (MSkip n)  MSkip ` {i. i  n}"
  by (induct i arbitrary: n) (auto dest: set_mp[OF RPDi_Test(1)] simp: le_Suc_eq split: nat.splits)

lemma RPDs_MSkip: "RPDs (MSkip n)  MSkip ` {i. i  n}"
  unfolding RPDs_def using RPDi_MSkip by auto

lemma RPDi_Plus: "RPDi i (MPlus r s)  {MPlus r s}  RPDi i r  RPDi i s"
  by (induct i arbitrary: r s) auto

lemma RPDi_Suc_RPD_Plus:
  "RPDi (Suc i) r  RPDs (MPlus r s)"
  "RPDi (Suc i) s  RPDs (MPlus r s)"
  unfolding RPDs_def by (force intro!: exI[of _ "Suc i"])+

lemma RPDs_Plus: "RPDs (MPlus r s)  {MPlus r s}  RPDs r  RPDs s"
  unfolding RPDs_def using RPDi_Plus[of _ r s] by auto

lemma RPDi_Times:
  "RPDi i (MTimes r s)  {MTimes r s}  MTimesL r (j  i. RPDi j s)  (j  i. RPDi j r)"
proof (induct i arbitrary: r s)
  case (Suc i)
  show ?case
    by (fastforce simp: MTimesL_def dest: bspec[of _ _ "Suc _"] dest!: set_mp[OF Suc])
qed simp

lemma RPDs_Times: "RPDs (MTimes r s)  {MTimes r s}  MTimesL r (RPDs s)  RPDs r"
  unfolding RPDs_def using RPDi_Times[of _ r s] by (force simp: MTimesL_def)

lemma RPDi_Star: "j  i  RPDi j (MStar r)  {MStar r}  MTimesL (MStar r) (j  i. RPDi j r)"
proof (induct i arbitrary: j r)
  case (Suc i)
  from Suc(2) show ?case
    by (auto 0 5 simp: MTimesL_def image_iff le_Suc_eq
        dest: bspec[of _ _ "Suc 0"] bspec[of _ _ "Suc _"] set_mp[OF Suc(1)] dest!: set_mp[OF RPDi_Times])
qed simp

lemma RPDs_Star: "RPDs (MStar r)  {MStar r}  MTimesL (MStar r) (RPDs r)"
  unfolding RPDs_def using RPDi_Star[OF order_refl, of _ r] by (force simp: MTimesL_def)

lemma finite_RPDs: "finite (RPDs r)"
proof (induct r)
  case MSkip
  then show ?case by (intro finite_subset[OF RPDs_MSkip]) simp
next
  case (MTestPos φ)
  then show ?case by (intro finite_subset[OF RPDs_Test(2)]) simp
next
  case (MTestNeg φ)
  then show ?case by (intro finite_subset[OF RPDs_Test(3)]) simp
next
  case (MPlus r s)
  then show ?case by (intro finite_subset[OF RPDs_Plus]) simp
next
  case (MTimes r s)
  then show ?case by (intro finite_subset[OF RPDs_Times]) (simp add: MTimesL_def)
next
  case (MStar r)
  then show ?case by (intro finite_subset[OF RPDs_Star]) (simp add: MTimesL_def)
qed

context begin

private abbreviation (input) "addRPD r  λS. insert r S  Set.bind (insert r S) RPD"

private lemma mono_addRPD: "mono (addRPD r)"
  unfolding mono_def Set.bind_def by auto

private lemma RPDs_aux1: "lfp (addRPD r)  RPDs r"
  by (rule lfp_induct[OF mono_addRPD], auto intro: RPDs_refl RPDs_trans simp: Set.bind_def)

private lemma RPDs_aux2: "RPDi i r  lfp (addRPD r)"
proof (induct i)
  case 0
  then show ?case
    by (subst lfp_unfold[OF mono_addRPD]) auto
next
  case (Suc i)
  then show ?case
    by (subst lfp_unfold[OF mono_addRPD]) (auto simp: RPDi_Suc_alt simp del: RPDi.simps)
qed

lemma RPDs_alt: "RPDs r = lfp (addRPD r)"
  using RPDs_aux1 RPDs_aux2 by (force simp: RPDs_def)

lemma RPDs_code[code]:
  "RPDs r = saturate (addRPD r) {}"
  unfolding RPDs_alt saturate_def
  by (rule lfp_while[OF mono_addRPD _ finite_RPDs, of r]) (auto simp: RPDs_refl RPDs_trans Set.bind_def)

end

subsection ‹The executable monitor›

type_synonym ts = nat

type_synonym 'a mbuf2 = "'a table list × 'a table list"
type_synonym 'a mbufn = "'a table list list"
type_synonym 'a msaux = "(ts × 'a table) list"
type_synonym 'a muaux = "(ts × 'a table × 'a table) list"
type_synonym 'a mrδaux = "(ts × (mregex, 'a table) mapping) list"
type_synonym 'a mlδaux = "(ts × 'a table list × 'a table) list"

datatype mconstraint = MEq | MLess | MLessEq

record args =
  args_ivl :: 
  args_n :: nat
  args_L :: "nat set"
  args_R :: "nat set"
  args_pos :: bool

datatype (dead 'msaux, dead 'muaux) mformula =
  MRel "event_data table"
  | MPred Formula.name "Formula.trm list"
  | MLet Formula.name nat "('msaux, 'muaux) mformula" "('msaux, 'muaux) mformula"
  | MAnd "nat set" "('msaux, 'muaux) mformula" bool "nat set" "('msaux, 'muaux) mformula" "event_data mbuf2"
  | MAndAssign "('msaux, 'muaux) mformula" "nat × Formula.trm"
  | MAndRel "('msaux, 'muaux) mformula" "Formula.trm × bool × mconstraint × Formula.trm"
  | MAnds "nat set list" "nat set list" "('msaux, 'muaux) mformula list" "event_data mbufn"
  | MOr "('msaux, 'muaux) mformula" "('msaux, 'muaux) mformula" "event_data mbuf2"
  | MNeg "('msaux, 'muaux) mformula"
  | MExists "('msaux, 'muaux) mformula"
  | MAgg bool nat Formula.agg_op nat "Formula.trm" "('msaux, 'muaux) mformula"
  | MPrev  "('msaux, 'muaux) mformula" bool "event_data table list" "ts list"
  | MNext  "('msaux, 'muaux) mformula" bool "ts list"
  | MSince args "('msaux, 'muaux) mformula" "('msaux, 'muaux) mformula" "event_data mbuf2" "ts list" "'msaux"
  | MUntil args "('msaux, 'muaux) mformula" "('msaux, 'muaux) mformula" "event_data mbuf2" "ts list" "'muaux"
  | MMatchP  "mregex" "mregex list" "('msaux, 'muaux) mformula list" "event_data mbufn" "ts list" "event_data mrδaux"
  | MMatchF  "mregex" "mregex list" "('msaux, 'muaux) mformula list" "event_data mbufn" "ts list" "event_data mlδaux"

record ('msaux, 'muaux) mstate =
  mstate_i :: nat
  mstate_m :: "('msaux, 'muaux) mformula"
  mstate_n :: nat

fun eq_rel :: "nat  Formula.trm  Formula.trm  event_data table" where
  "eq_rel n (Formula.Const x) (Formula.Const y) = (if x = y then unit_table n else empty_table)"
| "eq_rel n (Formula.Var x) (Formula.Const y) = singleton_table n x y"
| "eq_rel n (Formula.Const x) (Formula.Var y) = singleton_table n y x"
| "eq_rel n _ _ = undefined"

lemma regex_atms_size: "x  regex.atms r  size x < regex.size_regex size r"
  by (induct r) auto

lemma atms_size:
  assumes "x  atms r"
  shows "size x < Regex.size_regex size r"
proof -
  { fix y assume "y  regex.atms r" "case y of formula.Neg z  x = z | _  False"
    then have "size x < Regex.size_regex size r"
      by (cases y rule: formula.exhaust) (auto dest: regex_atms_size)
  }
  with assms show ?thesis
    unfolding atms_def
    by (auto split: formula.splits dest: regex_atms_size)
qed

definition init_args :: "  nat  nat set  nat set  bool  args" where
  "init_args I n L R pos = args_ivl = I, args_n = n, args_L = L, args_R = R, args_pos = pos"

locale msaux =
  fixes valid_msaux :: "args  ts  'msaux  event_data msaux  bool"
    and init_msaux :: "args  'msaux"
    and add_new_ts_msaux :: "args  ts  'msaux  'msaux"
    and join_msaux :: "args  event_data table  'msaux  'msaux"
    and add_new_table_msaux :: "args  event_data table  'msaux  'msaux"
    and result_msaux :: "args  'msaux  event_data table"
  assumes valid_init_msaux: "L  R 
    valid_msaux (init_args I n L R pos) 0 (init_msaux (init_args I n L R pos)) []"
  assumes valid_add_new_ts_msaux: "valid_msaux args cur aux auxlist  nt  cur 
    valid_msaux args nt (add_new_ts_msaux args nt aux)
    (filter (λ(t, rel). enat (nt - t)  right (args_ivl args)) auxlist)"
  assumes valid_join_msaux: "valid_msaux args cur aux auxlist 
    table (args_n args) (args_L args) rel1 
    valid_msaux args cur (join_msaux args rel1 aux)
    (map (λ(t, rel). (t, join rel (args_pos args) rel1)) auxlist)"
  assumes valid_add_new_table_msaux: "valid_msaux args cur aux auxlist 
    table (args_n args) (args_R args) rel2 
    valid_msaux args cur (add_new_table_msaux args rel2 aux)
    (case auxlist of
      [] => [(cur, rel2)]
    | ((t, y) # ts)  if t = cur then (t, y  rel2) # ts else (cur, rel2) # auxlist)"
    and valid_result_msaux: "valid_msaux args cur aux auxlist  result_msaux args aux =
    foldr (∪) [rel. (t, rel)  auxlist, left (args_ivl args)  cur - t] {}"

fun check_before :: "  ts  (ts × 'a × 'b)  bool" where
  "check_before I dt (t, a, b)  enat t + right I < enat dt"

fun proj_thd :: "('a × 'b × 'c)  'c" where
  "proj_thd (t, a1, a2) = a2"

definition update_until :: "args  event_data table  event_data table  ts  event_data muaux  event_data muaux" where
  "update_until args rel1 rel2 nt aux =
    (map (λx. case x of (t, a1, a2)  (t, if (args_pos args) then join a1 True rel1 else a1  rel1,
      if mem (nt - t) (args_ivl args) then a2  join rel2 (args_pos args) a1 else a2)) aux) @
    [(nt, rel1, if left (args_ivl args) = 0 then rel2 else empty_table)]"

lemma map_proj_thd_update_until: "map proj_thd (takeWhile (check_before (args_ivl args) nt) auxlist) =
  map proj_thd (takeWhile (check_before (args_ivl args) nt) (update_until args rel1 rel2 nt auxlist))"
proof (induction auxlist)
  case Nil
  then show ?case by (simp add: update_until_def)
next
  case (Cons a auxlist)
  then show ?case
    by (cases "right (args_ivl args)") (auto simp add: update_until_def split: if_splits prod.splits)
qed

fun eval_until :: "  ts  event_data muaux  event_data table list × event_data muaux" where
  "eval_until I nt [] = ([], [])"
| "eval_until I nt ((t, a1, a2) # aux) = (if t + right I < nt then
    (let (xs, aux) = eval_until I nt aux in (a2 # xs, aux)) else ([], (t, a1, a2) # aux))"

lemma eval_until_length:
  "eval_until I nt auxlist = (res, auxlist')  length auxlist = length res + length auxlist'"
  by (induction I nt auxlist arbitrary: res auxlist' rule: eval_until.induct)
    (auto split: if_splits prod.splits)

lemma eval_until_res: "eval_until I nt auxlist = (res, auxlist') 
  res = map proj_thd (takeWhile (check_before I nt) auxlist)"
  by (induction I nt auxlist arbitrary: res auxlist' rule: eval_until.induct)
    (auto split: prod.splits)

lemma eval_until_auxlist': "eval_until I nt auxlist = (res, auxlist') 
  auxlist' = drop (length res) auxlist"
  by (induction I nt auxlist arbitrary: res auxlist' rule: eval_until.induct)
    (auto split: if_splits prod.splits)

locale muaux =
  fixes valid_muaux :: "args  ts  'muaux  event_data muaux  bool"
    and init_muaux :: "args  'muaux"
    and add_new_muaux :: "args  event_data table  event_data table  ts  'muaux  'muaux"
    and length_muaux :: "args  'muaux  nat"
    and eval_muaux :: "args  ts  'muaux  event_data table list × 'muaux"
  assumes valid_init_muaux: "L  R 
    valid_muaux (init_args I n L R pos) 0 (init_muaux (init_args I n L R pos)) []"
  assumes valid_add_new_muaux: "valid_muaux args cur aux auxlist 
    table (args_n args) (args_L args) rel1 
    table (args_n args) (args_R args) rel2 
    nt  cur 
    valid_muaux args nt (add_new_muaux args rel1 rel2 nt aux)
    (update_until args rel1 rel2 nt auxlist)"
  assumes valid_length_muaux: "valid_muaux args cur aux auxlist  length_muaux args aux = length auxlist"
  assumes valid_eval_muaux: "valid_muaux args cur aux auxlist  nt  cur 
    eval_muaux args nt aux = (res, aux')  eval_until (args_ivl args) nt auxlist = (res', auxlist') 
    res = res'  valid_muaux args cur aux' auxlist'"

locale maux = msaux valid_msaux init_msaux add_new_ts_msaux join_msaux add_new_table_msaux result_msaux +
  muaux valid_muaux init_muaux add_new_muaux length_muaux eval_muaux
  for valid_msaux :: "args  ts  'msaux  event_data msaux  bool"
    and init_msaux :: "args  'msaux"
    and add_new_ts_msaux :: "args  ts  'msaux  'msaux"
    and join_msaux :: "args  event_data table  'msaux  'msaux"
    and add_new_table_msaux :: "args  event_data table  'msaux  'msaux"
    and result_msaux :: "args  'msaux  event_data table"
    and valid_muaux :: "args  ts  'muaux  event_data muaux  bool"
    and init_muaux :: "args  'muaux"
    and add_new_muaux :: "args  event_data table  event_data table  ts  'muaux  'muaux"
    and length_muaux :: "args  'muaux  nat"
    and eval_muaux :: "args  nat  'muaux  event_data table list × 'muaux"

fun split_assignment :: "nat set  Formula.formula  nat × Formula.trm" where
  "split_assignment X (Formula.Eq t1 t2) = (case (t1, t2) of
      (Formula.Var x, Formula.Var y)  if x  X then (y, t1) else (x, t2)
    | (Formula.Var x, _)  (x, t2)
    | (_, Formula.Var y)  (y, t1))"
| "split_assignment _ _ = undefined"

fun split_constraint :: "Formula.formula  Formula.trm × bool × mconstraint × Formula.trm" where
  "split_constraint (Formula.Eq t1 t2) = (t1, True, MEq, t2)"
| "split_constraint (Formula.Less t1 t2) = (t1, True, MLess, t2)"
| "split_constraint (Formula.LessEq t1 t2) = (t1, True, MLessEq, t2)"
| "split_constraint (Formula.Neg (Formula.Eq t1 t2)) = (t1, False, MEq, t2)"
| "split_constraint (Formula.Neg (Formula.Less t1 t2)) = (t1, False, MLess, t2)"
| "split_constraint (Formula.Neg (Formula.LessEq t1 t2)) = (t1, False, MLessEq, t2)"
| "split_constraint _ = undefined"

function (in maux) (sequential) minit0 :: "nat  Formula.formula  ('msaux, 'muaux) mformula" where
  "minit0 n (Formula.Neg φ) = (if fv φ = {} then MNeg (minit0 n φ) else MRel empty_table)"
| "minit0 n (Formula.Eq t1 t2) = MRel (eq_rel n t1 t2)"
| "minit0 n (Formula.Pred e ts) = MPred e ts"
| "minit0 n (Formula.Let p φ ψ) = MLet p (Formula.nfv φ) (minit0 (Formula.nfv φ) φ) (minit0 n ψ)"
| "minit0 n (Formula.Or φ ψ) = MOr (minit0 n φ) (minit0 n ψ) ([], [])"
| "minit0 n (Formula.And φ ψ) = (if safe_assignment (fv φ) ψ then
      MAndAssign (minit0 n φ) (split_assignment (fv φ) ψ)
    else if safe_formula ψ then
      MAnd (fv φ) (minit0 n φ) True (fv ψ) (minit0 n ψ) ([], [])
    else if is_constraint ψ then
      MAndRel (minit0 n φ) (split_constraint ψ)
    else (case ψ of Formula.Neg ψ 
      MAnd (fv φ) (minit0 n φ) False (fv ψ) (minit0 n ψ) ([], [])))"
| "minit0 n (Formula.Ands l) = (let (pos, neg) = partition safe_formula l in
    let mpos = map (minit0 n) pos in
    let mneg = map (minit0 n) (map remove_neg neg) in
    let vpos = map fv pos in
    let vneg = map fv neg in
    MAnds vpos vneg (mpos @ mneg) (replicate (length l) []))"
| "minit0 n (Formula.Exists φ) = MExists (minit0 (Suc n) φ)"
| "minit0 n (Formula.Agg y ω b f φ) = MAgg (fv φ  {0..<b}) y ω b f (minit0 (b + n) φ)"
| "minit0 n (Formula.Prev I φ) = MPrev I (minit0 n φ) True [] []"
| "minit0 n (Formula.Next I φ) = MNext I (minit0 n φ) True []"
| "minit0 n (Formula.Since φ I ψ) = (if safe_formula φ
    then MSince (init_args I n (Formula.fv φ) (Formula.fv ψ) True) (minit0 n φ) (minit0 n ψ) ([], []) [] (init_msaux (init_args I n (Formula.fv φ) (Formula.fv ψ) True))
    else (case φ of
      Formula.Neg φ  MSince (init_args I n (Formula.fv φ) (Formula.fv ψ) False) (minit0 n φ) (minit0 n ψ) ([], []) [] (init_msaux (init_args I n (Formula.fv φ) (Formula.fv ψ) False))
    | _  undefined))"
| "minit0 n (Formula.Until φ I ψ) = (if safe_formula φ
    then MUntil (init_args I n (Formula.fv φ) (Formula.fv ψ) True) (minit0 n φ) (minit0 n ψ) ([], []) [] (init_muaux (init_args I n (Formula.fv φ) (Formula.fv ψ) True))
    else (case φ of
      Formula.Neg φ  MUntil (init_args I n (Formula.fv φ) (Formula.fv ψ) False) (minit0 n φ) (minit0 n ψ) ([], []) [] (init_muaux (init_args I n (Formula.fv φ) (Formula.fv ψ) False))
    | _  undefined))"
| "minit0 n (Formula.MatchP I r) =
    (let (mr, φs) = to_mregex r
    in MMatchP I mr (sorted_list_of_set (RPDs mr)) (map (minit0 n) φs) (replicate (length φs) []) [] [])"
| "minit0 n (Formula.MatchF I r) =
    (let (mr, φs) = to_mregex r
    in MMatchF I mr (sorted_list_of_set (LPDs mr)) (map (minit0 n) φs) (replicate (length φs) []) [] [])"
| "minit0 n _ = undefined"
  by pat_completeness auto
termination (in maux)
  by (relation "measure (λ(_, φ). size φ)")
    (auto simp: less_Suc_eq_le size_list_estimation' size_remove_neg
      dest!: to_mregex_ok[OF sym] atms_size)

definition (in maux) minit :: "Formula.formula  ('msaux, 'muaux) mstate" where
  "minit φ = (let n = Formula.nfv φ in mstate_i = 0, mstate_m = minit0 n φ, mstate_n = n)"

definition (in maux) minit_safe where
  "minit_safe φ = (if mmonitorable_exec φ then minit φ else undefined)"

fun mprev_next :: "  event_data table list  ts list  event_data table list × event_data table list × ts list" where
  "mprev_next I [] ts = ([], [], ts)"
| "mprev_next I xs [] = ([], xs, [])"
| "mprev_next I xs [t] = ([], xs, [t])"
| "mprev_next I (x # xs) (t # t' # ts) = (let (ys, zs) = mprev_next I xs (t' # ts)
    in ((if mem (t' - t) I then x else empty_table) # ys, zs))"

fun mbuf2_add :: "event_data table list  event_data table list  event_data mbuf2  event_data mbuf2" where
  "mbuf2_add xs' ys' (xs, ys) = (xs @ xs', ys @ ys')"

fun mbuf2_take :: "(event_data table  event_data table  'b)  event_data mbuf2  'b list × event_data mbuf2" where
  "mbuf2_take f (x # xs, y # ys) = (let (zs, buf) = mbuf2_take f (xs, ys) in (f x y # zs, buf))"
| "mbuf2_take f (xs, ys) = ([], (xs, ys))"

fun mbuf2t_take :: "(event_data table  event_data table  ts  'b  'b)  'b 
    event_data mbuf2  ts list  'b × event_data mbuf2 × ts list" where
  "mbuf2t_take f z (x # xs, y # ys) (t # ts) = mbuf2t_take f (f x y t z) (xs, ys) ts"
| "mbuf2t_take f z (xs, ys) ts = (z, (xs, ys), ts)"

lemma size_list_length_diff1: "xs  []  []  set xs 
  size_list (λxs. length xs - Suc 0) xs < size_list length xs"
proof (induct xs)
  case (Cons x xs)
  then show ?case
    by (cases xs) auto
qed simp

fun mbufn_add :: "event_data table list list  event_data mbufn  event_data mbufn" where
  "mbufn_add xs' xs = List.map2 (@) xs xs'"

function mbufn_take :: "(event_data table list  'b  'b)  'b  event_data mbufn  'b × event_data mbufn" where
  "mbufn_take f z buf = (if buf = []  []  set buf then (z, buf)
    else mbufn_take f (f (map hd buf) z) (map tl buf))"
  by pat_completeness auto
termination by (relation "measure (λ(_, _, buf). size_list length buf)")
    (auto simp: comp_def Suc_le_eq size_list_length_diff1)

fun mbufnt_take :: "(event_data table list  ts  'b  'b) 
    'b  event_data mbufn  ts list  'b × event_data mbufn × ts list" where
  "mbufnt_take f z buf ts =
    (if []  set buf  ts = [] then (z, buf, ts)
    else mbufnt_take f (f (map hd buf) (hd ts) z) (map tl buf) (tl ts))"

fun match :: "Formula.trm list  event_data list  (nat  event_data) 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 meval_trm :: "Formula.trm  event_data tuple  event_data" where
  "meval_trm (Formula.Var x) v = the (v ! x)"
| "meval_trm (Formula.Const x) v = x"
| "meval_trm (Formula.Plus x y) v = meval_trm x v + meval_trm y v"
| "meval_trm (Formula.Minus x y) v = meval_trm x v - meval_trm y v"
| "meval_trm (Formula.UMinus x) v = - meval_trm x v"
| "meval_trm (Formula.Mult x y) v = meval_trm x v * meval_trm y v"
| "meval_trm (Formula.Div x y) v = meval_trm x v div meval_trm y v"
| "meval_trm (Formula.Mod x y) v = meval_trm x v mod meval_trm y v"
| "meval_trm (Formula.F2i x) v = EInt (integer_of_event_data (meval_trm x v))"
| "meval_trm (Formula.I2f x) v = EFloat (double_of_event_data (meval_trm x v))"

definition eval_agg :: "nat  bool  nat  Formula.agg_op  nat  Formula.trm 
  event_data table  event_data table" where
  "eval_agg n g0 y ω b f rel = (if g0  rel = empty_table
    then singleton_table n y (eval_agg_op ω {})
    else (λk.
      let group = Set.filter (λx. drop b x = k) rel;
        M = (λy. (y, ecard (Set.filter (λx. meval_trm f x = y) group))) ` meval_trm f ` group
      in k[y:=Some (eval_agg_op ω M)]) ` (drop b) ` rel)"

definition (in maux) update_since :: "args  event_data table  event_data table  ts 
  'msaux  event_data table × 'msaux" where
  "update_since args rel1 rel2 nt aux =
    (let aux0 = join_msaux args rel1 (add_new_ts_msaux args nt aux);
         aux' = add_new_table_msaux args rel2 aux0
    in (result_msaux args aux', aux'))"

definition "lookup = Mapping.lookup_default empty_table"

fun ε_lax where
  "ε_lax guard φs (MSkip n) = (if n = 0 then guard else empty_table)"
| "ε_lax guard φs (MTestPos i) = join guard True (φs ! i)"
| "ε_lax guard φs (MTestNeg i) = join guard False (φs ! i)"
| "ε_lax guard φs (MPlus r s) = ε_lax guard φs r  ε_lax guard φs s"
| "ε_lax guard φs (MTimes r s) = join (ε_lax guard φs r) True (ε_lax guard φs s)"
| "ε_lax guard φs (MStar r) = guard"

fun rε_strict where
  "rε_strict n φs (MSkip m) = (if m = 0 then unit_table n else empty_table)"
| "rε_strict n φs (MTestPos i) = φs ! i"
| "rε_strict n φs (MTestNeg i) = (if φs ! i = empty_table then unit_table n else empty_table)"
| "rε_strict n φs (MPlus r s) = rε_strict n φs r  rε_strict n φs s"
| "rε_strict n φs (MTimes r s) = ε_lax (rε_strict n φs r) φs s"
| "rε_strict n φs (MStar r) = unit_table n"

fun lε_strict where
  "lε_strict n φs (MSkip m) = (if m = 0 then unit_table n else empty_table)"
| "lε_strict n φs (MTestPos i) = φs ! i"
| "lε_strict n φs (MTestNeg i) = (if φs ! i = empty_table then unit_table n else empty_table)"
| "lε_strict n φs (MPlus r s) = lε_strict n φs r  lε_strict n φs s"
| "lε_strict n φs (MTimes r s) = ε_lax (lε_strict n φs s) φs r"
| "lε_strict n φs (MStar r) = unit_table n"

fun  :: "(mregex  mregex)  (mregex, 'a table) mapping  'a table list  mregex  'a table"  where
  " κ X φs (MSkip n) = (case n of 0  empty_table | Suc m  lookup X (κ (MSkip m)))"
| " κ X φs (MTestPos i) = empty_table"
| " κ X φs (MTestNeg i) = empty_table"
| " κ X φs (MPlus r s) =  κ X φs r   κ X φs s"
| " κ X φs (MTimes r s) =  (λt. κ (MTimes r t)) X φs s  ε_lax ( κ X φs r) φs s"
| " κ X φs (MStar r) =  (λt. κ (MTimes (MStar r) t)) X φs r"

fun  :: "(mregex  mregex)  (mregex, 'a table) mapping  'a table list  mregex  'a table" where
  " κ X φs (MSkip n) = (case n of 0  empty_table | Suc m  lookup X (κ (MSkip m)))"
| " κ X φs (MTestPos i) = empty_table"
| " κ X φs (MTestNeg i) = empty_table"
| " κ X φs (MPlus r s) =  κ X φs r   κ X φs s"
| " κ X φs (MTimes r s) =  (λt. κ (MTimes t s)) X φs r  ε_lax ( κ X φs s) φs r"
| " κ X φs (MStar r) =  (λt. κ (MTimes t (MStar r))) X φs r"

lift_definition mrtabulate :: "mregex list  (mregex  'b table)  (mregex, 'b table) mapping"
  is "λks f. (map_of (List.map_filter (λk. let fk = f k in if fk = empty_table then None else Some (k, fk)) ks))" .

lemma lookup_tabulate:
  "distinct xs  lookup (mrtabulate xs f) x = (if x  set xs then f x else empty_table)"
  unfolding lookup_default_def lookup_def
  by transfer (auto simp: Let_def map_filter_def map_of_eq_None_iff o_def image_image dest!: map_of_SomeD
      split: if_splits option.splits)

definition update_matchP :: "nat    mregex  mregex list  event_data table list  ts 
  event_data mrδaux  event_data table × event_data mrδaux" where
  "update_matchP n I mr mrs rels nt aux =
    (let aux = (case [(t, mrtabulate mrs (λmr.
         id rel rels mr  (if t = nt then rε_strict n rels mr else {}))).
      (t, rel)  aux, enat (nt - t)  right I]
      of []  [(nt, mrtabulate mrs (rε_strict n rels))]
      | x # aux'  (if fst x = nt then x # aux'
                     else (nt, mrtabulate mrs (rε_strict n rels)) # x # aux'))
    in (foldr (∪) [lookup rel mr. (t, rel)  aux, left I  nt - t] {}, aux))"

definition update_matchF_base where
  "update_matchF_base n I mr mrs rels nt =
     (let X = mrtabulate mrs (lε_strict n rels)
     in ([(nt, rels, if left I = 0 then lookup X mr else empty_table)], X))"

definition update_matchF_step where
  "update_matchF_step I mr mrs nt = (λ(t, rels', rel) (aux', X).
     (let Y = mrtabulate mrs ( id X rels')
     in ((t, rels', if mem (nt - t) I then rel  lookup Y mr else rel) # aux', Y)))"

definition update_matchF :: "nat    mregex  mregex list  event_data table list  ts  event_data mlδaux  event_data mlδaux" where
  "update_matchF n I mr mrs rels nt aux =
     fst (foldr (update_matchF_step I mr mrs nt) aux (update_matchF_base n I mr mrs rels nt))"

fun eval_matchF :: "  mregex  ts  event_data mlδaux  event_data table list × event_data mlδaux" where
  "eval_matchF I mr nt [] = ([], [])"
| "eval_matchF I mr nt ((t, rels, rel) # aux) = (if t + right I < nt then
    (let (xs, aux) = eval_matchF I mr nt aux in (rel # xs, aux)) else ([], (t, rels, rel) # aux))"

primrec map_split where
  "map_split f [] = ([], [])"
| "map_split f (x # xs) =
    (let (y, z) = f x; (ys, zs) = map_split f xs
    in (y # ys, z # zs))"

fun eval_assignment :: "nat × Formula.trm  event_data tuple  event_data tuple" where
  "eval_assignment (x, t) y = (y[x:=Some (meval_trm t y)])"

fun eval_constraint0 :: "mconstraint  event_data  event_data  bool" where
  "eval_constraint0 MEq x y = (x = y)"
| "eval_constraint0 MLess x y = (x < y)"
| "eval_constraint0 MLessEq x y = (x  y)"

fun eval_constraint :: "Formula.trm × bool × mconstraint × Formula.trm  event_data tuple  bool" where
  "eval_constraint (t1, p, c, t2) x = (eval_constraint0 c (meval_trm t1 x) (meval_trm t2 x) = p)"

primrec (in maux) meval :: "nat  ts  Formula.database  ('msaux, 'muaux) mformula 
    event_data table list × ('msaux, 'muaux) mformula" where
  "meval n t db (MRel rel) = ([rel], MRel rel)"
| "meval n t db (MPred e ts) = (map (λX. (λf. Table.tabulate f 0 n) ` Option.these
    (match ts ` X)) (case Mapping.lookup db e of None  [{}] | Some xs  xs), MPred e ts)"
| "meval n t db (MLet p m φ ψ) =
    (let (xs, φ) = meval m t db φ; (ys, ψ) = meval n t (Mapping.update p (map (image (map the)) xs) db) ψ
    in (ys, MLet p m φ ψ))"
| "meval n t db (MAnd A_φ φ pos A_ψ ψ buf) =
    (let (xs, φ) = meval n t db φ; (ys, ψ) = meval n t db ψ;
      (zs, buf) = mbuf2_take (λr1 r2. bin_join n A_φ r1 pos A_ψ r2) (mbuf2_add xs ys buf)
    in (zs, MAnd A_φ φ pos A_ψ ψ buf))"
| "meval n t db (MAndAssign φ conf) =
    (let (xs, φ) = meval n t db φ in (map (λr. eval_assignment conf ` r) xs, MAndAssign φ conf))"
| "meval n t db (MAndRel φ conf) =
    (let (xs, φ) = meval n t db φ in (map (Set.filter (eval_constraint conf)) xs, MAndRel φ conf))"
| "meval n t db (MAnds A_pos A_neg L buf) =
    (let R = map (meval n t db) L in
    let buf = mbufn_add (map fst R) buf in
    let (zs, buf) = mbufn_take (λxs zs. zs @ [mmulti_join n A_pos A_neg xs]) [] buf in
    (zs, MAnds A_pos A_neg (map snd R) buf))"
| "meval n t db (MOr φ ψ buf) =
    (let (xs, φ) = meval n t db φ; (ys, ψ) = meval n t db ψ;
      (zs, buf) = mbuf2_take (λr1 r2. r1  r2) (mbuf2_add xs ys buf)
    in (zs, MOr φ ψ buf))"
| "meval n t db (MNeg φ) =
    (let (xs, φ) = meval n t db φ in (map (λr. (if r = empty_table then unit_table n else empty_table)) xs, MNeg φ))"
| "meval n t db (MExists φ) =
    (let (xs, φ) = meval (Suc n) t db φ in (map (λr. tl ` r) xs, MExists φ))"
| "meval n t db (MAgg g0 y ω b f φ) =
    (let (xs, φ) = meval (b + n) t db φ in (map (eval_agg n g0 y ω b f) xs, MAgg g0 y ω b f φ))"
| "meval n t db (MPrev I φ first buf nts) =
    (let (xs, φ) = meval n t db φ;
      (zs, buf, nts) = mprev_next I (buf @ xs) (nts @ [t])
    in (if first then empty_table # zs else zs, MPrev I φ False buf nts))"
| "meval n t db (MNext I φ first nts) =
    (let (xs, φ) = meval n t db φ;
      (xs, first) = (case (xs, first) of (_ # xs, True)  (xs, False) | a  a);
      (zs, _, nts) = mprev_next I xs (nts @ [t])
    in (zs, MNext I φ first nts))"
| "meval n t db (MSince args φ ψ buf nts aux) =
    (let (xs, φ) = meval n t db φ; (ys, ψ) = meval n t db ψ;
      ((zs, aux), buf, nts) = mbuf2t_take (λr1 r2 t (zs, aux).
        let (z, aux) = update_since args r1 r2 t aux
        in (zs @ [z], aux)) ([], aux) (mbuf2_add xs ys buf) (nts @ [t])
    in (zs, MSince args φ ψ buf nts aux))"
| "meval n t db (MUntil args φ ψ buf nts aux) =
    (let (xs, φ) = meval n t db φ; (ys, ψ) = meval n t db ψ;
      (aux, buf, nts) = mbuf2t_take (add_new_muaux args) aux (mbuf2_add xs ys buf) (nts @ [t]);
      (zs, aux) = eval_muaux args (case nts of []  t | nt # _  nt) aux
    in (zs, MUntil args φ ψ buf nts aux))"
| "meval n t db (MMatchP I mr mrs φs buf nts aux) =
    (let (xss, φs) = map_split id (map (meval n t db) φs);
      ((zs, aux), buf, nts) = mbufnt_take (λrels t (zs, aux).
        let (z, aux) = update_matchP n I mr mrs rels t aux
        in (zs @ [z], aux)) ([], aux) (mbufn_add xss buf) (nts @ [t])
    in (zs, MMatchP I mr mrs φs buf nts aux))"
| "meval n t db (MMatchF I mr mrs φs buf nts aux) =
    (let (xss, φs) = map_split id (map (meval n t db) φs);
      (aux, buf, nts) = mbufnt_take (update_matchF n I mr mrs) aux (mbufn_add xss buf) (nts @ [t]);
      (zs, aux) = eval_matchF I mr (case nts of []  t | nt # _  nt) aux
    in (zs, MMatchF I mr mrs φs buf nts aux))"

definition (in maux) mstep :: "Formula.database × ts  ('msaux, 'muaux) mstate  (nat × event_data table) list × ('msaux, 'muaux) mstate" where
  "mstep tdb st =
     (let (xs, m) = meval (mstate_n st) (snd tdb) (fst tdb) (mstate_m st)
     in (List.enumerate (mstate_i st) xs,
      mstate_i = mstate_i st + length xs, mstate_m = m, mstate_n = mstate_n st))"

subsection ‹Verdict delay›

context fixes σ :: Formula.trace begin

fun progress :: "(Formula.name  nat)  Formula.formula  nat  nat" where
  "progress P (Formula.Pred e ts) j = (case P e of None  j | Some k  k)"
| "progress P (Formula.Let p φ ψ) j = progress (P(p  progress P φ j)) ψ j"
| "progress P (Formula.Eq t1 t2) j = j"
| "progress P (Formula.Less t1 t2) j = j"
| "progress P (Formula.LessEq t1 t2) j = j"
| "progress P (Formula.Neg φ) j = progress P φ j"
| "progress P (Formula.Or φ ψ) j = min (progress P φ j) (progress P ψ j)"
| "progress P (Formula.And φ ψ) j = min (progress P φ j) (progress P ψ j)"
| "progress P (Formula.Ands l) j = (if l = [] then j else Min (set (map (λφ. progress P φ j) l)))"
| "progress P (Formula.Exists φ) j = progress P φ j"
| "progress P (Formula.Agg y ω b f φ) j = progress P φ j"
| "progress P (Formula.Prev I φ) j = (if j = 0 then 0 else min (Suc (progress P φ j)) j)"
| "progress P (Formula.Next I φ) j = progress P φ j - 1"
| "progress P (Formula.Since φ I ψ) j = min (progress P φ j) (progress P ψ j)"
| "progress P (Formula.Until φ I ψ) j =
    Inf {i. k. k < j  k  min (progress P φ j) (progress P ψ j)  τ σ i + right I  τ σ k}"
| "progress P (Formula.MatchP I r) j = min_regex_default (progress P) r j"
| "progress P (Formula.MatchF I r) j =
    Inf {i. k. k < j  k  min_regex_default (progress P) r j  τ σ i + right I  τ σ k}"

definition "progress_regex P = min_regex_default (progress P)"

declare progress.simps[simp del]
lemmas progress_simps[simp] = progress.simps[folded progress_regex_def[THEN fun_cong, THEN fun_cong]]

end

definition "pred_mapping Q = pred_fun (λ_. True) (pred_option Q)"
definition "rel_mapping Q = rel_fun (=) (rel_option Q)"

lemma pred_mapping_alt: "pred_mapping Q P = (p  dom P. Q (the (P p)))"
  unfolding pred_mapping_def pred_fun_def option.pred_set dom_def
  by (force split: option.splits)

lemma rel_mapping_alt: "rel_mapping Q P P' = (dom P = dom P'  (p  dom P. Q (the (P p)) (the (P' p))))"
  unfolding rel_mapping_def rel_fun_def rel_option_iff dom_def
  by (force split: option.splits)

lemma rel_mapping_map_upd[simp]: "Q x y  rel_mapping Q P P'  rel_mapping Q (P(p  x)) (P'(p  y))"
  by (auto simp: rel_mapping_alt)

lemma pred_mapping_map_upd[simp]: "Q x  pred_mapping Q P  pred_mapping Q (P(p  x))"
  by (auto simp: pred_mapping_alt)

lemma pred_mapping_empty[simp]: "pred_mapping Q Map.empty"
  by (auto simp: pred_mapping_alt)

lemma pred_mapping_mono: "pred_mapping Q P  Q  R  pred_mapping R P"
  by (auto simp: pred_mapping_alt)

lemma pred_mapping_mono_strong: "pred_mapping Q P 
  (p. p  dom P  Q (the (P p))  R (the (P p)))  pred_mapping R P"
  by (auto simp: pred_mapping_alt)

lemma progress_mono_gen: "j  j'  rel_mapping (≤) P P'  progress σ P φ j  progress σ P' φ j'"
proof (induction φ arbitrary: P P')
  case (Pred e ts)
  then show ?case
    by (force simp: rel_mapping_alt dom_def split: option.splits)
next
  case (Ands l)
  then show ?case
    by (auto simp: image_iff intro!: Min.coboundedI[THEN order_trans])
next
  case (Until φ I ψ)
  from Until(1,2)[of P P'] Until.prems show ?case
    by (cases "right I")
      (auto dest: trans_le_add1[OF τ_mono] intro!: cInf_superset_mono)
next
  case (MatchF I r)
  from MatchF(1)[of _ P P'] MatchF.prems show ?case
    by (cases "right I"; cases "regex.atms r = {}")
      (auto 0 3 simp: Min_le_iff progress_regex_def dest: trans_le_add1[OF τ_mono]
        intro!: cInf_superset_mono elim!: less_le_trans order_trans)
qed (force simp: Min_le_iff progress_regex_def split: option.splits)+

lemma rel_mapping_reflp: "reflp Q  rel_mapping Q P P"
  by (auto simp: rel_mapping_alt reflp_def)

lemmas progress_mono = progress_mono_gen[OF _ rel_mapping_reflp[unfolded reflp_def], simplified]

lemma progress_le_gen: "pred_mapping (λx. x  j) P  progress σ P φ j  j"
proof (induction φ arbitrary: P)
  case (Pred e ts)
  then show ?case
    by (auto simp: pred_mapping_alt dom_def split: option.splits)
next
  case (Ands l)
  then show ?case
    by (auto simp: image_iff intro!: Min.coboundedI[where a="progress σ P (hd l) j", THEN order_trans])
next
  case (Until φ I ψ)
  then show ?case
    by (cases "right I")
      (auto intro: trans_le_add1[OF τ_mono] intro!: cInf_lower)
next
  case (MatchF I r)
  then show ?case
    by (cases "right I")
      (auto intro: trans_le_add1[OF τ_mono] intro!: cInf_lower)
qed (force simp: Min_le_iff progress_regex_def split: option.splits)+

lemma progress_le: "progress σ Map.empty φ j  j"
  using progress_le_gen[of _ Map.empty] by auto

lemma progress_0_gen[simp]:
  "pred_mapping (λx. x = 0) P  progress σ P φ 0 = 0"
  using progress_le_gen[of 0 P] by auto

lemma progress_0[simp]:
  "progress σ Map.empty φ 0 = 0"
  by (auto simp: pred_mapping_alt)

definition max_mapping :: "('b  'a option)  ('b  'a option)  'b  ('a :: linorder) option" where
  "max_mapping P P' x = (case (P x, P' x) of
    (None, None)  None
  | (Some x, None)  None
  | (None, Some x)  None
  | (Some x, Some y)  Some (max x y))"

definition Max_mapping :: "('b  'a option) set  'b  ('a :: linorder) option" where
  "Max_mapping Ps x = (if (P  Ps. P x  None) then Some (Max ((λP. the (P x)) ` Ps)) else None)"

lemma dom_max_mapping[simp]: "dom (max_mapping P1 P2) = dom P1  dom P2"
  unfolding max_mapping_def by (auto split: option.splits)

lemma dom_Max_mapping[simp]: "dom (Max_mapping X) = (P  X. dom P)"
  unfolding Max_mapping_def by (auto split: if_splits)

lemma Max_mapping_coboundedI:
  assumes "finite X" "Q  X. dom Q = dom P" "P  X"
  shows "rel_mapping (≤) P (Max_mapping X)"
  unfolding rel_mapping_alt
proof (intro conjI ballI)
  from assms(3) have "X  {}" by auto
  then show "dom P = dom (Max_mapping X)" using assms(2) by auto
next
  fix p
  assume "p  dom P"
  with assms show "the (P p)  the (Max_mapping X p)"
    by (force simp add: Max_mapping_def intro!: Max.coboundedI imageI)
qed

lemma rel_mapping_trans: "P OO Q  R 
  rel_mapping P P1 P2  rel_mapping Q P2 P3  rel_mapping R P1 P3"
  by (force simp: rel_mapping_alt dom_def set_eq_iff)

abbreviation range_mapping :: "nat  nat  ('b  nat option)  bool" where
  "range_mapping i j P  pred_mapping (λx. i  x  x  j) P"

lemma range_mapping_relax:
  "range_mapping i j P  i'  i  j'  j  range_mapping i' j' P"
  by (auto simp: pred_mapping_alt dom_def set_eq_iff max_mapping_def split: option.splits)

lemma range_mapping_max_mapping[simp]:
  "range_mapping i j1 P1  range_mapping i j2 P2  range_mapping i (max j1 j2) (max_mapping P1 P2)"
  by (auto simp: pred_mapping_alt dom_def set_eq_iff max_mapping_def split: option.splits)

lemma range_mapping_Max_mapping[simp]:
  "finite X  X  {}  xX. range_mapping i (j x) (P x)  range_mapping i (Max (j ` X)) (Max_mapping (P ` X))"
  by (force simp: pred_mapping_alt Max_mapping_def dom_def image_iff
      intro!: Max_ge_iff[THEN iffD2] split: if_splits)

lemma pred_mapping_le:
  "pred_mapping ((≤) i) P1  rel_mapping (≤) P1 P2  pred_mapping ((≤) (i :: nat)) P2"
  by (force simp: rel_mapping_alt pred_mapping_alt dom_def set_eq_iff)

lemma pred_mapping_le':
  "pred_mapping ((≤) j) P1  i  j  rel_mapping (≤) P1 P2  pred_mapping ((≤) (i :: nat)) P2"
  by (force simp: rel_mapping_alt pred_mapping_alt dom_def set_eq_iff)

lemma max_mapping_cobounded1: "dom P1  dom P2  rel_mapping (≤) P1 (max_mapping P1 P2)"
  unfolding max_mapping_def rel_mapping_alt by (auto simp: dom_def split: option.splits)

lemma max_mapping_cobounded2: "dom P2  dom P1  rel_mapping (≤) P2 (max_mapping P1 P2)"
  unfolding max_mapping_def rel_mapping_alt by (auto simp: dom_def split: option.splits)

lemma max_mapping_fun_upd2[simp]:
  "(max_mapping P1 (P2(p := y)))(p  x) = (max_mapping P1 P2)(p  x)"
  by (auto simp: max_mapping_def)

lemma rel_mapping_max_mapping_fun_upd: "dom P2  dom P1  p  dom P2  the (P2 p)  y 
  rel_mapping (≤) P2 ((max_mapping P1 P2)(p  y))"
  by (auto simp: rel_mapping_alt max_mapping_def split: option.splits)

lemma progress_ge_gen: "Formula.future_bounded φ 
   P j. dom P = S  range_mapping i j P  i  progress σ P φ j"
proof (induction φ arbitrary: i S)
  case (Pred e ts)
  then show ?case
    by (intro exI[of _ "λe. if e  S then Some i else None"])
      (auto split: option.splits if_splits simp: rel_mapping_alt pred_mapping_alt dom_def)
next
  case (Let p φ ψ)
  from Let.prems obtain P2 j2 where P2: "dom P2 = insert p S" "range_mapping i j2 P2"
    "i  progress σ P2 ψ j2"
    by (atomize_elim, intro Let(2)) (force simp: pred_mapping_alt rel_mapping_alt dom_def)+
  from Let.prems obtain P1 j1 where P1: "dom P1 = S" "range_mapping (the (P2 p)) j1 P1"
    "the (P2 p)  progress σ P1 φ j1"
    by (atomize_elim, intro Let(1)) auto
  let ?P12 = "max_mapping P1 P2"
  from P1 P2 have le1: "progress σ P1 φ j1  progress σ (?P12(p := P1 p)) φ (max j1 j2)"
    by (intro progress_mono_gen) (auto simp: rel_mapping_alt max_mapping_def)
  from P1 P2 have le2: "progress σ P2 ψ j2  progress σ (?P12(p  progress σ P1 φ j1)) ψ (max j1 j2)"
    by (intro progress_mono_gen) (auto simp: rel_mapping_alt max_mapping_def)
  show ?case
    unfolding progress.simps
  proof (intro exI[of _ "?P12(p := P1 p)"] exI[of _ "max j1 j2"] conjI)
    show "dom (?P12(p := P1 p)) = S"
      using P1 P2 by (auto simp: dom_def max_mapping_def)
  next
    show "range_mapping i (max j1 j2) (?P12(p := P1 p))"
      using P1 P2 by (force simp add: pred_mapping_alt dom_def max_mapping_def split: option.splits)
  next
    have "i  progress σ P2 ψ j2" by fact
    also have "...  progress σ (?P12(p  progress σ P1 φ j1)) ψ (max j1 j2)"
      using le2 by blast
    also have "...  progress σ ((?P12(p := P1 p))(pprogress σ (?P12(p := P1 p)) φ (max j1 j2))) ψ (max j1 j2)"
      by (auto intro!: progress_mono_gen simp: le1 rel_mapping_alt)
    finally show "i  ..." .
  qed
next
  case (Eq _ _)
  then show ?case
    by (intro exI[of _ "λe. if e  S then Some i else None"]) (auto split: if_splits simp: pred_mapping_alt)
next
  case (Less _ _)
  then show ?case
    by (intro exI[of _ "λe. if e  S then Some i else None"]) (auto split: if_splits simp: pred_mapping_alt)
next
  case (LessEq _ _)
  then show ?case
    by (intro exI[of _ "λe. if e  S then Some i else None"]) (auto split: if_splits simp: pred_mapping_alt)
next
  case (Or φ1 φ2)
  from Or(3) obtain P1 j1 where P1: "dom P1 = S" "range_mapping i j1 P1"  "i  progress σ P1 φ1 j1"
    using Or(1)[of S i] by auto
  moreover
  from Or(3) obtain P2 j2 where P2: "dom P2 = S" "range_mapping i j2 P2" "i  progress σ P2 φ2 j2"
    using Or(2)[of S i] by auto
  ultimately have "i  progress σ (max_mapping P1 P2) (Formula.Or φ1 φ2) (max j1 j2)"
    by (auto 0 3 elim!: order.trans[OF _ progress_mono_gen] intro: max_mapping_cobounded1 max_mapping_cobounded2)
  with P1 P2 show ?case by (intro exI[of _ "max_mapping P1 P2"] exI[of _ "max j1 j2"]) auto
next
  case (And φ1 φ2)
  from And(3) obtain P1 j1 where P1: "dom P1 = S" "range_mapping i j1 P1"  "i  progress σ P1 φ1 j1"
    using And(1)[of S i] by auto
  moreover
  from And(3) obtain P2 j2 where P2: "dom P2 = S" "range_mapping i j2 P2" "i  progress σ P2 φ2 j2"
    using And(2)[of S i] by auto
  ultimately have "i  progress σ (max_mapping P1 P2) (Formula.And φ1 φ2) (max j1 j2)"
    by (auto 0 3 elim!: order.trans[OF _ progress_mono_gen] intro: max_mapping_cobounded1 max_mapping_cobounded2)
  with P1 P2 show ?case by (intro exI[of _ "max_mapping P1 P2"] exI[of _ "max j1 j2"]) auto
next
  case (Ands l)
  show ?case proof (cases "l = []")
    case True
    then show ?thesis
      by (intro exI[of _ "λe. if e  S then Some i else None"])
        (auto split: if_splits simp: pred_mapping_alt)
  next
    case False
    then obtain φ where "φ  set l" by (cases l) auto
    from Ands.prems have "φset l. Formula.future_bounded φ"
      by (simp add: list.pred_set)
    { fix φ
      assume "φ  set l"
      with Ands.prems obtain P j where "dom P = S" "range_mapping i j P" "i  progress σ P φ j"
        by (atomize_elim, intro Ands(1)[of φ S i]) (auto simp: list.pred_set)
      then have "Pj. dom (fst Pj) = S  range_mapping i (snd Pj) (fst Pj)  i  progress σ (fst Pj) φ (snd Pj)"
        (is "Pj. ?P Pj")
        by (intro exI[of _ "(P, j)"]) auto
    }
    then have "φset l. Pj. dom (fst Pj) = S  range_mapping i (snd Pj) (fst Pj)  i  progress σ (fst Pj) φ (snd Pj)"
      (is "φset l. Pj. ?P Pj φ")
      by blast
    with Ands(1) Ands.prems False have "Pjf. φset l. ?P (Pjf φ) φ"
      by (auto simp: Ball_def intro: choice)
    then obtain Pjf where Pjf: "φset l. ?P (Pjf φ) φ" ..
    define Pf where "Pf = fst o Pjf"
    define jf where "jf = snd o Pjf"
    have *: "dom (Pf φ) = S" "range_mapping i (jf φ) (Pf φ)" "i  progress σ (Pf φ) φ (jf φ)"
      if "φ  set l" for φ
      using Pjf[THEN bspec, OF that] unfolding Pf_def jf_def by auto
    with False show ?thesis
      unfolding progress.simps eq_False[THEN iffD2, OF False] if_False
      by ((subst Min_ge_iff; simp add: False),
          intro exI[where x="MAX φset l. jf φ"] exI[where x="Max_mapping (Pf ` set l)"]
          conjI ballI order.trans[OF *(3) progress_mono_gen] Max_mapping_coboundedI)
        (auto simp: False *[OF φ  set l] φ  set l)
  qed
next
  case (Exists φ)
  then show ?case by simp
next
  case (Prev I φ)
  then obtain P j where "dom P = S" "range_mapping i j P" "i  progress σ P φ j"
    by (atomize_elim, intro Prev(1)) (auto simp: pred_mapping_alt dom_def)
  with Prev(2) have
    "dom P = S  range_mapping i (max i j) P  i  progress σ P (formula.Prev I φ) (max i j)"
    by (auto simp: le_Suc_eq max_def pred_mapping_alt split: if_splits
        elim: order.trans[OF _ progress_mono])
  then show ?case by blast
next
  case (Next I φ)
  then obtain P j where "dom P = S" "range_mapping (Suc i) j P" "Suc i  progress σ P φ j"
    by (atomize_elim, intro Next(1)) (auto simp: pred_mapping_alt dom_def)
  then show ?case
    by (intro exI[of _ P] exI[of _ j]) (auto 0 3 simp: pred_mapping_alt dom_def)
next
  case (Since φ1 I φ2)
  from Since(3) obtain P1 j1 where P1: "dom P1 = S" "range_mapping i j1 P1"  "i  progress σ P1 φ1 j1"
    using Since(1)[of S i] by auto
  moreover
  from Since(3) obtain P2 j2 where P2: "dom P2 = S" "range_mapping i j2 P2" "i  progress σ P2 φ2 j2"
    using Since(2)[of S i] by auto
  ultimately have "i  progress σ (max_mapping P1 P2) (Formula.Since φ1 I φ2) (max j1 j2)"
    by (auto elim!: order.trans[OF _ progress_mono_gen] simp: max_mapping_cobounded1 max_mapping_cobounded2)
  with P1 P2 show ?case by (intro exI[of _ "max_mapping P1 P2"] exI[of _ "max j1 j2"])
      (auto elim!: pred_mapping_le intro: max_mapping_cobounded1)
next
  case (Until φ1 I φ2)
  from Until.prems obtain b where [simp]: "right I = enat b"
    by (cases "right I") (auto)
  obtain i' where "i < i'" and "τ σ i + b + 1  τ σ i'"
    using ex_le_τ[where x="τ σ i + b + 1"] by (auto simp add: less_eq_Suc_le)
  then have 1: "τ σ i + b < τ σ i'" by simp
  from Until.prems obtain P1 j1 where P1: "dom P1 = S" "range_mapping (Suc i') j1 P1" "Suc i'  progress σ P1 φ1 j1"
    by (atomize_elim, intro Until(1)) (auto simp: pred_mapping_alt dom_def)
  from Until.prems obtain P2 j2 where P2: "dom P2 = S" "range_mapping (Suc i') j2 P2" "Suc i'  progress σ P2 φ2 j2"
    by (atomize_elim, intro Until(2)) (auto simp: pred_mapping_alt dom_def)
  let ?P12 = "max_mapping P1 P2"
  have "i  progress σ ?P12 (Formula.Until φ1 I φ2) (max j1 j2)"
    unfolding progress.simps
  proof (intro cInf_greatest, goal_cases nonempty greatest)
    case nonempty
    then show ?case
      by (auto simp: trans_le_add1[OF τ_mono] intro!: exI[of _ "max j1 j2"])
  next
    case (greatest x)
    from P1(2,3) have "i' < j1"
      by (auto simp: less_eq_Suc_le intro!: progress_le_gen elim!: order.trans pred_mapping_mono_strong)
    then have "i' < max j1 j2" by simp
    have "progress σ P1 φ1 j1  progress σ ?P12 φ1 (max j1 j2)"
      using P1(1) P2(1) by (auto intro!: progress_mono_gen max_mapping_cobounded1)
    moreover have "progress σ P2 φ2 j2  progress σ ?P12 φ2 (max j1 j2)"
      using P1(1) P2(1) by (auto intro!: progress_mono_gen max_mapping_cobounded2)
    ultimately have "i'  min (progress σ ?P12 φ1 (max j1 j2)) (progress σ ?P12 φ2 (max j1 j2))"
      using P1(3) P2(3) by simp
    with greatest i' < max j1 j2 have "τ σ i'  τ σ x + b"
      by simp
    with 1 have "τ σ i < τ σ x" by simp
    then show ?case by (auto dest!: less_τD)
  qed
  with P1 P2 i < i' show ?case
    by (intro exI[of _ "max_mapping P1 P2"] exI[of _ "max j1 j2"]) (auto simp: range_mapping_relax)
next
  case (MatchP I r)
  then show ?case
  proof (cases "regex.atms r = {}")
    case True
    with MatchP.prems show ?thesis
      unfolding progress.simps
      by (intro exI[of _ "λe. if e  S then Some i else None"] exI[of _ i])
        (auto split: if_splits simp: pred_mapping_alt regex.pred_set)
  next
    case False
    define pick where "pick = (λφ. SOME Pj. dom (fst Pj) = S  range_mapping i (snd Pj) (fst Pj) 
       i  progress σ (fst Pj) φ (snd Pj))"
    let ?pickP = "fst o pick" let ?pickj = "snd o pick"
    from MatchP have pick: "φ  regex.atms r  dom (?pickP φ) = S 
      range_mapping i (?pickj φ) (?pickP φ)  i  progress σ (?pickP φ) φ (?pickj φ)" for φ
      unfolding pick_def o_def future_bounded.simps regex.pred_set
      by (intro someI_ex[where P = "λPj. dom (fst Pj) = S  range_mapping i (snd Pj) (fst Pj) 
         i  progress σ (fst Pj) φ (snd Pj)"]) auto
    with False show ?thesis
      unfolding progress.simps
      by (intro exI[of _ "Max_mapping (?pickP ` regex.atms r)"] exI[of _ "Max (?pickj ` regex.atms r)"])
        (auto simp: Max_mapping_coboundedI
          order_trans[OF pick[THEN conjunct2, THEN conjunct2] progress_mono_gen])
  qed
next
  case (MatchF I r)
  from MatchF.prems obtain b where [simp]: "right I = enat b"
    by auto
  obtain i' where i': "i < i'" "τ σ i + b + 1  τ σ i'"
    using ex_le_τ[where x="τ σ i + b + 1"] by (auto simp add: less_eq_Suc_le)
  then have 1: "τ σ i + b < τ σ i'" by simp
  have ix: "i  x" if "τ σ i'  b + τ σ x" "b + τ σ i < τ σ i'" for x
    using less_τD[of σ i] that less_le_trans by fastforce
  show ?case
  proof (cases "regex.atms r = {}")
    case True
    with MatchF.prems i' ix show ?thesis
      unfolding progress.simps
      by (intro exI[of _ "λe. if e  S then Some (Suc i') else None"] exI[of _ "Suc i'"])
        (auto split: if_splits simp: pred_mapping_alt regex.pred_set add.commute less_Suc_eq
          intro!: cInf_greatest dest!: spec[of _ i'] less_imp_le[THEN τ_mono, of _ i' σ])
  next
    case False
    then obtain φ where φ: "φ  regex.atms r" by auto
    define pick where "pick = (λφ. SOME Pj. dom (fst Pj) = S  range_mapping (Suc i') (snd Pj) (fst Pj) 
      Suc i'  progress σ (fst Pj) φ (snd Pj))"
    define pickP where "pickP = fst o pick" define pickj where "pickj = snd o pick"
    from MatchF have pick: "φ  regex.atms r  dom (pickP φ) = S 
    range_mapping (Suc i') (pickj φ) (pickP φ)  Suc i'  progress σ (pickP φ) φ (pickj φ)" for φ
      unfolding pick_def o_def future_bounded.simps regex.pred_set pickj_def pickP_def
      by (intro someI_ex[where P = "λPj. dom (fst Pj) = S  range_mapping (Suc i') (snd Pj) (fst Pj) 
       Suc i'  progress σ (fst Pj) φ (snd Pj)"]) auto
    let ?P = "Max_mapping (pickP ` regex.atms r)" let ?j = "Max (pickj ` regex.atms r)"
    from pick[OF φ] False φ have "Suc i'  ?j"
      by (intro order_trans[OF pick[THEN conjunct2, THEN conjunct2], OF φ] order_trans[OF progress_le_gen])
        (auto simp: Max_ge_iff dest: range_mapping_relax[of _ _ _ 0, OF _ _ order_refl, simplified])
    moreover
    note i' 1 ix
    moreover
    from MatchF.prems have "Regex.pred_regex Formula.future_bounded r"
      by auto
    ultimately show ?thesis using τ_mono[of _ ?j σ] less_τD[of σ i] pick False
      by (intro exI[of _ "?j"]  exI[of _ "?P"])
        (auto 0 3 intro!: cInf_greatest
          order_trans[OF le_SucI[OF order_refl] order_trans[OF pick[THEN conjunct2, THEN conjunct2] progress_mono_gen]]
          range_mapping_Max_mapping[OF _ _ ballI[OF range_mapping_relax[of "Suc i'" _ _ i, OF _ _ order_refl]]]
          simp: ac_simps Suc_le_eq trans_le_add2 Max_mapping_coboundedI progress_regex_def
          dest: spec[of _ "i'"] spec[of _ ?j])
  qed
qed (auto split: option.splits)

lemma progress_ge: "Formula.future_bounded φ  j. i  progress σ Map.empty φ j"
  using progress_ge_gen[of φ "{}" i σ]
  by auto

lemma cInf_restrict_nat:
  fixes x :: nat
  assumes "x  A"
  shows "Inf A = Inf {y  A. y  x}"
  using assms by (auto intro!: antisym intro: cInf_greatest cInf_lower Inf_nat_def1)

lemma progress_time_conv:
  assumes "i<j. τ σ i = τ σ' i"
  shows "progress σ P φ j = progress σ' P φ j"
  using assms proof (induction φ arbitrary: P)
  case (Until φ1 I φ2)
  have *: "i  j - 1  i < j" if "j  0" for i
    using that by auto
  with Until show ?case
  proof (cases "right I")
    case (enat b)
    then show ?thesis
    proof (cases "j")
      case (Suc n)
      with enat * Until show ?thesis
        using τ_mono[THEN trans_le_add1]
        by (auto 8 0
            intro!: box_equals[OF arg_cong[where f=Inf]
              cInf_restrict_nat[symmetric, where x=n] cInf_restrict_nat[symmetric, where x=n]])
    qed simp
  qed simp
next
  case (MatchF I r)
  have *: "i  j - 1  i < j" if "j  0" for i
    using that by auto
  with MatchF show ?case using τ_mono[THEN trans_le_add1]
    by (cases "right I"; cases j)
      ((auto 6 0 simp: progress_le_gen progress_regex_def intro!: box_equals[OF arg_cong[where f=Inf]
            cInf_restrict_nat[symmetric, where x="j-1"] cInf_restrict_nat[symmetric, where x="j-1"]]) [])+
qed (auto simp: progress_regex_def)

lemma Inf_UNIV_nat: "(Inf UNIV :: nat) = 0"
  by (simp add: cInf_eq_minimum)

lemma progress_prefix_conv:
  assumes "prefix_of π σ" and "prefix_of π σ'"
  shows "progress σ P φ (plen π) = progress σ' P φ (plen π)"
  using assms by (auto intro: progress_time_conv τ_prefix_conv)

lemma bounded_rtranclp_mono:
  fixes n :: "'x :: linorder"
  assumes "i j. R i j  j < n  S i j" "i j. R i j  i  j"
  shows "rtranclp R i j  j < n  rtranclp S i j"
proof (induct rule: rtranclp_induct)
  case (step y z)
  then show ?case
    using assms(1,2)[of y z]
    by (auto elim!: rtrancl_into_rtrancl[to_pred, rotated])
qed auto

lemma sat_prefix_conv_gen:
  assumes "prefix_of π σ" and "prefix_of π σ'"
  shows "i < progress σ P φ (plen π)  dom V = dom V'  dom P = dom V 
    pred_mapping (λx. x  plen π) P 
    (p i φ. p  dom V  i < the (P p)  the (V p) i = the (V' p) i) 
    Formula.sat σ V v i φ  Formula.sat σ' V' v i φ"
proof (induction φ arbitrary: P V V' v i)
  case (Pred e ts)
  from Pred.prems(1,4) have "i < plen π"
    by (blast intro!: order.strict_trans2 progress_le_gen)
  show ?case proof (cases "V e")
    case None
    then have "V' e = None" using dom V = dom V' by auto
    with None Γ_prefix_conv[OF assms(1,2) i < plen π] show ?thesis by simp
  next
    case (Some a)
    obtain a' where "V' e = Some a'" using Some dom V = dom V' by auto
    then have "i < the (P e)"
      using Pred.prems(1-3) by (auto split: option.splits)
    then have "the (V e) i = the (V' e) i"
      using Some by (intro Pred.prems(5)) (simp_all add: domI)
    with Some V' e = Some a' show ?thesis by simp
  qed
next
  case (Let p φ ψ)
  let ?V = "λV σ. (V(p  λi. {v. length v = Formula.nfv φ  Formula.sat σ V v i φ}))"
  show ?case unfolding sat.simps proof (rule Let.IH(2))
    from Let.prems show "i < progress σ (P(p  progress σ P φ (plen π))) ψ (plen π)"
      by simp
    from Let.prems show "dom (?V V σ) = dom (?V V' σ')"
      by simp
    from Let.prems show "dom (P(p  progress σ P φ (plen π))) = dom (?V V σ)"
      by simp
    from Let.prems show "pred_mapping (λx. x  plen π) (P(p  progress σ P φ (plen π)))"
      by (auto intro!: pred_mapping_map_upd elim!: progress_le_gen)
  next
    fix p' i φ'
    assume 1: "p'  dom (?V V σ)" and 2: "i < the ((P(p  progress σ P φ (plen π))) p')"
    show "the (?V V σ p') i = the (?V V' σ' p') i" proof (cases "p' = p")
      case True
      with Let 2 show ?thesis by auto
    next
      case False
      with 1 2 show ?thesis by (auto intro!: Let.prems(5))
    qed
  qed
next
  case (Eq t1 t2)
  show ?case by simp
next
  case (Neg φ)
  then show ?case by simp
next
  case (Or φ1 φ2)
  then show ?case by auto
next
  case (Ands l)
  from Ands.prems have "φset l. i < progress σ P φ (plen π)"
    by (cases l) simp_all
  with Ands show ?case unfolding sat_Ands by blast
next
  case (Exists φ)
  then show ?case by simp
next
  case (Prev I φ)
  with τ_prefix_conv[OF assms(1,2)] show ?case
    by (cases i) (auto split: if_splits)
next
  case (Next I φ)
  then have "Suc i < plen π"
    by (auto intro: order.strict_trans2[OF _ progress_le_gen[of _ P σ φ]])
  with Next.prems τ_prefix_conv[OF assms(1,2)] show ?case
    unfolding sat.simps
    by (intro conj_cong Next) auto
next
  case (Since φ1 I φ2)
  then have "i < plen π"
    by (auto elim!: order.strict_trans2[OF _ progress_le_gen])
  with Since.prems τ_prefix_conv[OF assms(1,2)] show ?case
    unfolding sat.simps
    by (intro conj_cong ex_cong ball_cong Since) auto
next
  case (Until φ1 I φ2)
  from Until.prems obtain b where right[simp]: "right I = enat b"
    by (cases "right I") (auto simp add: Inf_UNIV_nat)
  from Until.prems obtain j where "τ σ i + b + 1  τ σ j"
    "j  progress σ P φ1 (plen π)" "j  progress σ P φ2 (plen π)"
    by atomize_elim (auto 0 4 simp add: less_eq_Suc_le not_le intro: Suc_leI dest: spec[of _ "i"]
        dest!: le_cInf_iff[THEN iffD1, rotated -1])
  then have 1: "k < progress σ P φ1 (plen π)" and 2: "k < progress σ P φ2 (plen π)"
    if "τ σ k  τ σ i + b" for k
    using that by (fastforce elim!: order.strict_trans2[rotated] intro: less_τD[of σ])+
  have 3: "k < plen π" if "τ σ k  τ σ i + b" for k
    using 1[OF that] Until(6) by (auto simp only: less_eq_Suc_le order.trans[OF _ progress_le_gen])

  from Until.prems have "i < progress σ' P (Formula.Until φ1 I φ2) (plen π)"
    unfolding progress_prefix_conv[OF assms(1,2)] by blast
  then obtain j where "τ σ' i + b + 1  τ σ' j"
    "j  progress σ' P φ1 (plen π)" "j  progress σ' P φ2 (plen π)"
    by atomize_elim (auto 0 4 simp add: less_eq_Suc_le not_le intro: Suc_leI dest: spec[of _ "i"]
        dest!: le_cInf_iff[THEN iffD1, rotated -1])
  then have 11: "k < progress σ P φ1 (plen π)" and 21: "k < progress σ P φ2 (plen π)"
    if "τ σ' k  τ σ' i + b" for k
    unfolding progress_prefix_conv[OF assms(1,2)]
    using that by (fastforce elim!: order.strict_trans2[rotated] intro: less_τD[of σ'])+
  have 31: "k < plen π" if "τ σ' k  τ σ' i + b" for k
    using 11[OF that] Until(6) by (auto simp only: less_eq_Suc_le order.trans[OF _ progress_le_gen])
  show ?case unfolding sat.simps
  proof ((intro ex_cong iffI; elim conjE), goal_cases LR RL)
    case (LR j)
    with Until(1)[OF 1] Until(2)[OF 2] τ_prefix_conv[OF assms(1,2) 3] Until.prems show ?case
      by (auto 0 4 simp: le_diff_conv add.commute dest: less_imp_le order.trans[OF τ_mono, rotated])
  next
    case (RL j)
    with Until(1)[OF 11] Until(2)[OF 21] τ_prefix_conv[OF assms(1,2) 31] Until.prems show ?case
      by (auto 0 4 simp: le_diff_conv add.commute dest: less_imp_le order.trans[OF τ_mono, rotated])
  qed
next
  case (MatchP I r)
  then have "i < plen π"
    by (force simp: pred_mapping_alt elim!: order.strict_trans2[OF _ progress_le_gen])
  with MatchP.prems τ_prefix_conv[OF assms(1,2)] show ?case
    unfolding sat.simps
    by (intro ex_cong conj_cong match_cong_strong MatchP) (auto simp: progress_regex_def split: if_splits)
next
  case (MatchF I r)
  from MatchF.prems obtain b where right[simp]: "right I = enat b"
    by (cases "right I") (auto simp add: Inf_UNIV_nat)
  show ?case
  proof (cases "regex.atms r = {}")
    case True
    from MatchF.prems(1) obtain j where "τ σ i + b + 1  τ σ j" "j  plen π"
      by atomize_elim (auto 0 4 simp add: less_eq_Suc_le not_le dest!: le_cInf_iff[THEN iffD1, rotated -1])
    then have 1: "k < plen π" if "τ σ k  τ σ i + b" for k
      by (meson τ_mono discrete not_le order.strict_trans2 that)
    from MatchF.prems have "i < progress σ' P (Formula.MatchF I r) (plen π)"
      unfolding progress_prefix_conv[OF assms(1,2)] by blast
    then obtain j where "τ σ' i + b + 1  τ σ' j" "j  plen π"
      by atomize_elim (auto 0 4 simp add: less_eq_Suc_le not_le dest!: le_cInf_iff[THEN iffD1, rotated -1])
    then have 2: "k < plen π" if "τ σ' k  τ σ' i + b" for k
      by (meson τ_mono discrete not_le order.strict_trans2 that)
    from MatchF.prems(1,4) True show ?thesis
      unfolding sat.simps conj_commute[of "left I  _" "_  _"]
    proof (intro ex_cong conj_cong match_cong_strong, goal_cases left right sat)
      case (left j)
      then show ?case
        by (intro iffI)
          ((subst (1 2) τ_prefix_conv[OF assms(1,2) 1, symmetric]; auto elim: order.trans[OF τ_mono, rotated]),
            (subst (1 2) τ_prefix_conv[OF assms(1,2) 2]; auto elim: order.trans[OF τ_mono, rotated]))
    next
      case (right j)
      then show ?case
        by (intro iffI)
          ((subst (1 2) τ_prefix_conv[OF assms(1,2) 2, symmetric]; auto elim: order.trans[OF τ_mono, rotated]),
            (subst (1 2) τ_prefix_conv[OF assms(1,2) 2]; auto elim: order.trans[OF τ_mono, rotated]))
    qed auto
  next
    case False
    from MatchF.prems(1) False obtain j where "τ σ i + b + 1  τ σ j" "(xregex.atms r. j  progress σ P x (plen π))"
      by atomize_elim (auto 0 6 simp add: less_eq_Suc_le not_le progress_regex_def
        dest!: le_cInf_iff[THEN iffD1, rotated -1])
    then have 1: "φ  regex.atms r  k < progress σ P φ (plen π)" if "τ σ k  τ σ i + b" for k φ
      using that
      by (fastforce elim!: order.strict_trans2[rotated] intro: less_τD[of σ])
    then have 2: "k < plen π" if "τ σ k  τ σ i + b" "regex.atms r  {}" for k
      using that
      by (fastforce intro: order.strict_trans2[OF _ progress_le_gen[OF MatchF(5), of σ], of k])

    from MatchF.prems have "i < progress σ' P (Formula.MatchF I r) (plen π)"
      unfolding progress_prefix_conv[OF assms(1,2)] by blast
    with False obtain j where "τ σ' i + b + 1  τ σ' j" "(xregex.atms r. j  progress σ' P x (plen π))"
      by atomize_elim (auto 0 6 simp add: less_eq_Suc_le not_le progress_regex_def
        dest!: le_cInf_iff[THEN iffD1, rotated -1])
    then have 11:  "φ  regex.atms r  k < progress σ P φ (plen π)" if "τ σ' k  τ σ' i + b" for k φ
      using that using progress_prefix_conv[OF assms(1,2)]
      by (auto 0 3 elim!: order.strict_trans2[rotated] intro: less_τD[of σ'])
    have 21: "k < plen π" if "τ σ' k  τ σ' i + b" for k
      using 11[OF that(1)] False by (fastforce intro: order.strict_trans2[OF _ progress_le_gen[OF MatchF(5), of σ], of k])
    show ?thesis unfolding sat.simps conj_commute[of "left I  _" "_  _"]
    proof ((intro ex_cong conj_cong match_cong_strong MatchF(1)[OF _ _ MatchF(3-6)]; assumption?), goal_cases right left progress)
      case (right j)
      with False show ?case
        by (intro iffI)
          ((subst (1 2) τ_prefix_conv[OF assms(1,2) 2, symmetric]; auto elim: order.trans[OF τ_mono, rotated]),
            (subst (1 2) τ_prefix_conv[OF assms(1,2) 21]; auto elim: order.trans[OF τ_mono, rotated]))
    next
      case (left j)
      with False show ?case unfolding right enat_ord_code le_diff_conv add.commute[of b]
        by (intro iffI)
          ((subst (1 2) τ_prefix_conv[OF assms(1,2) 21, symmetric]; auto elim: order.trans[OF τ_mono, rotated]),
            (subst (1 2) τ_prefix_conv[OF assms(1,2) 21]; auto elim: order.trans[OF τ_mono, rotated]))
    next
      case (progress j k z)
      with False show ?case unfolding right enat_ord_code le_diff_conv add.commute[of b]
        by (elim 1[rotated])
          (subst (1 2) τ_prefix_conv[OF assms(1,2) 21]; auto elim!: order.trans[OF τ_mono, rotated])
    qed
  qed
qed auto

lemma sat_prefix_conv:
  assumes "prefix_of π σ" and "prefix_of π σ'"
  shows "i < progress σ Map.empty φ (plen π) 
    Formula.sat σ Map.empty v i φ  Formula.sat σ' Map.empty v i φ"
  by (erule sat_prefix_conv_gen[OF assms]) auto

lemma progress_remove_neg[simp]: "progress σ P (remove_neg φ) j = progress σ P φ j"
  by (cases φ) simp_all

lemma safe_progress_get_and: "safe_formula φ 
  Min ((λφ. progress σ P φ j) ` set (get_and_list φ)) = progress σ P φ j"
  by (induction φ rule: get_and_list.induct) auto

lemma progress_convert_multiway: "safe_formula φ  progress σ P (convert_multiway φ) j = progress σ P φ j"
proof (induction φ arbitrary: P rule: safe_formula_induct)
  case (And_safe φ ψ)
  let ?c = "convert_multiway (Formula.And φ ψ)"
  let ?cφ = "convert_multiway φ"
  let ?cψ = "convert_multiway ψ"
  have c_eq: "?c = Formula.Ands (get_and_list ?cφ @ get_and_list ?cψ)"
    using And_safe by simp
  from safe_formula φ have "safe_formula ?cφ" by (rule safe_convert_multiway)
  moreover from safe_formula ψ have "safe_formula ?cψ" by (rule safe_convert_multiway)
  ultimately show ?case
    unfolding c_eq
    using And_safe.IH
    by (auto simp: get_and_nonempty Min.union safe_progress_get_and)
next
  case (And_Not φ ψ)
  let ?c = "convert_multiway (Formula.And φ (Formula.Neg ψ))"
  let ?cφ = "convert_multiway φ"
  let ?cψ = "convert_multiway ψ"
  have c_eq: "?c = Formula.Ands (Formula.Neg ?cψ # get_and_list ?cφ)"
    using And_Not by simp
  from safe_formula φ have "safe_formula ?cφ" by (rule safe_convert_multiway)
  moreover from safe_formula ψ have "safe_formula ?cψ" by (rule safe_convert_multiway)
  ultimately show ?case
    unfolding c_eq
    using And_Not.IH
    by (auto simp: get_and_nonempty Min.union safe_progress_get_and)
next
  case (MatchP I r)
  from MatchP show ?case
    unfolding progress.simps regex.map convert_multiway.simps regex.set_map image_image
    by (intro if_cong arg_cong[of _ _ Min] image_cong)
      (auto 0 4 simp: atms_def elim!: disjE_Not2 dest: safe_regex_safe_formula)
next
  case (MatchF I r)
  from MatchF show ?case
    unfolding progress.simps regex.map convert_multiway.simps regex.set_map image_image
    by (intro if_cong arg_cong[of _ _ Min] arg_cong[of _ _ Inf] arg_cong[of _ _ "(≤) _"]
      image_cong Collect_cong all_cong1 imp_cong conj_cong image_cong)
      (auto 0 4 simp: atms_def elim!: disjE_Not2 dest: safe_regex_safe_formula)
qed auto


subsection ‹Specification›

definition pprogress :: "Formula.formula  Formula.prefix  nat" where
  "pprogress φ π = (THE n. σ. prefix_of π σ  progress σ Map.empty φ (plen π) = n)"

lemma pprogress_eq: "prefix_of π σ  pprogress φ π = progress σ Map.empty φ (plen π)"
  unfolding pprogress_def using progress_prefix_conv
  by blast

locale future_bounded_mfodl =
  fixes φ :: Formula.formula
  assumes future_bounded: "Formula.future_bounded φ"

sublocale future_bounded_mfodl  sliceable_timed_progress "Formula.nfv φ" "Formula.fv φ" "relevant_events φ"
  "λσ v i. Formula.sat σ Map.empty v i φ" "pprogress φ"
proof (unfold_locales, goal_cases)
  case (1 x)
  then show ?case by (simp add: fvi_less_nfv)
next
  case (2 v v' σ i)
  then show ?case by (simp cong: sat_fv_cong[rule_format])
next
  case (3 v S σ i)
  then show ?case
    using sat_slice_iff[symmetric] by simp
next
  case (4 π π')
  moreover obtain σ where "prefix_of π' σ"
    using ex_prefix_of ..
  moreover have "prefix_of π σ"
    using prefix_of_antimono[OF π  π' prefix_of π' σ] .
  ultimately show ?case
    by (simp add: pprogress_eq plen_mono progress_mono)
next
  case (5 σ x)
  obtain j where "x  progress σ Map.empty φ j"
    using future_bounded progress_ge by blast
  then have "x  pprogress φ (take_prefix j σ)"
    by (simp add: pprogress_eq[of _ σ])
  then show ?case by force
next
  case (6 π σ σ' i v)
  then have "i < progress σ Map.empty φ (plen π)"
    by (simp add: pprogress_eq)
  with 6 show ?case
    using sat_prefix_conv by blast
next
  case (7 π π')
  then have "plen π = plen π'"
    by transfer (simp add: list_eq_iff_nth_eq)
  moreover obtain σ σ' where "prefix_of π σ" "prefix_of π' σ'"
    using ex_prefix_of by blast+
  moreover have "i < plen π. τ σ i = τ σ' i"
    using 7 calculation
    by transfer (simp add: list_eq_iff_nth_eq)
  ultimately show ?case
    by (simp add: pprogress_eq progress_time_conv)
qed

locale verimon_spec =
  fixes φ :: Formula.formula
  assumes monitorable: "mmonitorable φ"

sublocale verimon_spec  future_bounded_mfodl
  using monitorable by unfold_locales (simp add: mmonitorable_def)


subsection ‹Correctness›

subsubsection ‹Invariants›

definition wf_mbuf2 :: "nat  nat  nat  (nat  event_data table  bool)  (nat  event_data table  bool) 
    event_data mbuf2  bool" where
  "wf_mbuf2 i ja jb P Q buf  i  ja  i  jb  (case buf of (xs, ys) 
    list_all2 P [i..<ja] xs  list_all2 Q [i..<jb] ys)"

inductive list_all3 :: "('a  'b  'c  bool)  'a list  'b list  'c list  bool" for P :: "('a  'b  'c  bool)" where
  "list_all3 P [] [] []"
| "P a1 a2 a3  list_all3 P q1 q2 q3  list_all3 P (a1 # q1) (a2 # q2) (a3 # q3)"

lemma list_all3_list_all2D: "list_all3 P xs ys zs 
  (length xs = length ys  list_all2 (case_prod P) (zip xs ys) zs)"
  by (induct xs ys zs rule: list_all3.induct) auto

lemma list_all2_list_all3I: "length xs = length ys  list_all2 (case_prod P) (zip xs ys) zs 
  list_all3 P xs ys zs"
  by (induct xs ys arbitrary: zs rule: list_induct2)
    (auto simp: list_all2_Cons1 intro: list_all3.intros)

lemma list_all3_list_all2_eq: "list_all3 P xs ys zs 
  (length xs = length ys  list_all2 (case_prod P) (zip xs ys) zs)"
  using list_all2_list_all3I list_all3_list_all2D by blast

lemma list_all3_mapD: "list_all3 P (map f xs) (map g ys) (map h zs) 
  list_all3 (λx y z. P (f x) (g y) (h z)) xs ys zs"
  by (induct "map f xs" "map g ys" "map h zs" arbitrary: xs ys zs rule: list_all3.induct)
    (auto intro: list_all3.intros)

lemma list_all3_mapI: "list_all3 (λx y z. P (f x) (g y) (h z)) xs ys zs 
  list_all3 P (map f xs) (map g ys) (map h zs)"
  by (induct xs ys zs rule: list_all3.induct)
    (auto intro: list_all3.intros)

lemma list_all3_map_iff: "list_all3 P (map f xs) (map g ys) (map h zs) 
  list_all3 (λx y z. P (f x) (g y) (h z)) xs ys zs"
  using list_all3_mapD list_all3_mapI by blast

lemmas list_all3_map =
  list_all3_map_iff[where g=id and h=id, unfolded list.map_id id_apply]
  list_all3_map_iff[where f=id and h=id, unfolded list.map_id id_apply]
  list_all3_map_iff[where f=id and g=id, unfolded list.map_id id_apply]

lemma list_all3_conv_all_nth:
  "list_all3 P xs ys zs =
  (length xs = length ys  length ys = length zs  (i < length xs. P (xs!i) (ys!i) (zs!i)))"
  by (auto simp add: list_all3_list_all2_eq list_all2_conv_all_nth)

lemma list_all3_refl [intro?]:
  "(x. x  set xs  P x x x)  list_all3 P xs xs xs"
  by (simp add: list_all3_conv_all_nth)

definition wf_mbufn :: "nat  nat list  (nat  event_data table  bool) list  event_data mbufn  bool" where
  "wf_mbufn i js Ps buf  list_all3 (λP j xs. i  j  list_all2 P [i..<j] xs) Ps js buf"

definition wf_mbuf2' :: "Formula.trace  _  _  nat  nat  event_data list set 
  Formula.formula  Formula.formula  event_data mbuf2  bool" where
  "wf_mbuf2' σ P V j n R φ ψ buf  wf_mbuf2 (min (progress σ P φ j) (progress σ P ψ j))
    (progress σ P φ j) (progress σ P ψ j)
    (λi. qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) i φ))
    (λi. qtable n (Formula.fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) i ψ)) buf"

definition wf_mbufn' :: "Formula.trace  _  _  nat  nat  event_data list set 
  Formula.formula Regex.regex  event_data mbufn  bool" where
  "wf_mbufn' σ  P V j n R r buf  (case to_mregex r of (mr, φs) 
    wf_mbufn (progress_regex σ P r j) (map (λφ. progress σ P φ j) φs)
    (map (λφ i. qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) i φ)) φs)
    buf)"

lemma wf_mbuf2'_UNIV_alt: "wf_mbuf2' σ P V j n UNIV φ ψ buf  (case buf of (xs, ys) 
  list_all2 (λi. wf_table n (Formula.fv φ) (λv. Formula.sat σ V (map the v) i φ))
    [min (progress σ P φ j) (progress σ P ψ j) ..< (progress σ P φ j)] xs 
  list_all2 (λi. wf_table n (Formula.fv ψ) (λv. Formula.sat σ V (map the v) i ψ))
    [min (progress σ P φ j) (progress σ P ψ j) ..< (progress σ P ψ j)] ys)"
  unfolding wf_mbuf2'_def wf_mbuf2_def
  by (simp add: mem_restr_UNIV[THEN eqTrueI, abs_def] split: prod.split)

definition wf_ts :: "Formula.trace  _  nat  Formula.formula  Formula.formula  ts list  bool" where
  "wf_ts σ P j φ ψ ts  list_all2 (λi t. t = τ σ i) [min (progress σ P φ j) (progress σ P ψ j)..<j] ts"

definition wf_ts_regex :: "Formula.trace  _  nat  Formula.formula Regex.regex  ts list  bool" where
  "wf_ts_regex σ P j r ts  list_all2 (λi t. t = τ σ i) [progress_regex σ P r j..<j] ts"

abbreviation "Sincep pos φ I ψ  Formula.Since (if pos then φ else Formula.Neg φ) I ψ"

definition (in msaux) wf_since_aux :: "Formula.trace  _  event_data list set  args 
  Formula.formula  Formula.formula  'msaux  nat  bool" where
  "wf_since_aux σ V R args φ ψ aux ne  Formula.fv φ  Formula.fv ψ  (cur auxlist. valid_msaux args cur aux auxlist 
    cur = (if ne = 0 then 0 else τ σ (ne - 1)) 
    sorted_wrt (λx y. fst x > fst y) auxlist 
    (t X. (t, X)  set auxlist  ne  0  t  τ σ (ne - 1)  τ σ (ne - 1) - t  right (args_ivl args)  (i. τ σ i = t) 
      qtable (args_n args) (Formula.fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) (ne-1) (Sincep (args_pos args) φ (point (τ σ (ne - 1) - t)) ψ)) X) 
    (t. ne  0  t  τ σ (ne - 1)  τ σ (ne - 1) - t  right (args_ivl args)  (i. τ σ i = t) 
      (X. (t, X)  set auxlist)))"

definition wf_matchP_aux :: "Formula.trace  _  nat  event_data list set 
      Formula.formula Regex.regex  event_data mrδaux  nat  bool" where
  "wf_matchP_aux σ V n R I r aux ne  sorted_wrt (λx y. fst x > fst y) aux 
    (t X. (t, X)  set aux  ne  0  t  τ σ (ne-1)  τ σ (ne-1) - t  right I  (i. τ σ i = t) 
      (case to_mregex r of (mr, φs) 
      (ms  RPDs mr. qtable n (Formula.fv_regex r) (mem_restr R) (λv. Formula.sat σ V (map the v) (ne-1)
         (Formula.MatchP (point (τ σ (ne-1) - t)) (from_mregex ms φs)))
         (lookup X ms)))) 
    (t. ne  0  t  τ σ (ne-1)  τ σ (ne-1) - t  right I  (i. τ σ i = t) 
      (X. (t, X)  set aux))"

lemma qtable_mem_restr_UNIV: "qtable n A(mem_restr UNIV) Q X = wf_table n A Q X"
  unfolding qtable_def by auto

lemma (in msaux) wf_since_aux_UNIV_alt:
  "wf_since_aux σ V UNIV args φ ψ aux ne  Formula.fv φ  Formula.fv ψ  (cur auxlist. valid_msaux args cur aux auxlist 
    cur = (if ne = 0 then 0 else τ σ (ne - 1)) 
    sorted_wrt (λx y. fst x > fst y) auxlist 
    (t X. (t, X)  set auxlist  ne  0  t  τ σ (ne - 1)  τ σ (ne - 1) - t  right (args_ivl args)  (i. τ σ i = t) 
      wf_table (args_n args) (Formula.fv ψ)
          (λv. Formula.sat σ V (map the v) (ne - 1) (Sincep (args_pos args) φ (point (τ σ (ne - 1) - t)) ψ)) X) 
    (t. ne  0  t  τ σ (ne - 1)  τ σ (ne - 1) - t  right (args_ivl args)  (i. τ σ i = t) 
      (X. (t, X)  set auxlist)))"
  unfolding wf_since_aux_def qtable_mem_restr_UNIV ..

definition wf_until_auxlist :: "Formula.trace  _  nat  event_data list set  bool 
    Formula.formula    Formula.formula  event_data muaux  nat  bool" where
  "wf_until_auxlist σ V n R pos φ I ψ auxlist ne 
    list_all2 (λx i. case x of (t, r1, r2)  t = τ σ i 
      qtable n (Formula.fv φ) (mem_restr R) (λv. if pos then (k{i..<ne+length auxlist}. Formula.sat σ V (map the v) k φ)
          else (k{i..<ne+length auxlist}. Formula.sat σ V (map the v) k φ)) r1 
      qtable n (Formula.fv ψ) (mem_restr R) (λv. (j. i  j  j < ne + length auxlist  mem (τ σ j - τ σ i) I 
          Formula.sat σ V (map the v) j ψ 
          (k{i..<j}. if pos then Formula.sat σ V (map the v) k φ else ¬ Formula.sat σ V (map the v) k φ))) r2)
      auxlist [ne..<ne+length auxlist]"

definition (in muaux) wf_until_aux :: "Formula.trace  _  event_data list set  args 
  Formula.formula  Formula.formula  'muaux  nat  bool" where
  "wf_until_aux σ V R args φ ψ aux ne  Formula.fv φ  Formula.fv ψ 
    (cur auxlist. valid_muaux args cur aux auxlist 
      cur = (if ne + length auxlist = 0 then 0 else τ σ (ne + length auxlist - 1)) 
      wf_until_auxlist σ V (args_n args) R (args_pos args) φ (args_ivl args) ψ auxlist ne)"

lemma (in muaux) wf_until_aux_UNIV_alt:
  "wf_until_aux σ V UNIV args φ ψ aux ne  Formula.fv φ  Formula.fv ψ 
    (cur auxlist. valid_muaux args cur aux auxlist 
      cur = (if ne + length auxlist = 0 then 0 else τ σ (ne + length auxlist - 1)) 
      list_all2 (λx i. case x of (t, r1, r2)  t = τ σ i 
      wf_table (args_n args) (Formula.fv φ) (λv. if (args_pos args)
          then (k{i..<ne+length auxlist}. Formula.sat σ V (map the v) k φ)
          else (k{i..<ne+length auxlist}. Formula.sat σ V (map the v) k φ)) r1 
      wf_table (args_n args) (Formula.fv ψ) (λv. j. i  j  j < ne + length auxlist  mem (τ σ j - τ σ i) (args_ivl args) 
          Formula.sat σ V (map the v) j ψ 
          (k{i..<j}. if (args_pos args) then Formula.sat σ V (map the v) k φ else ¬ Formula.sat σ V (map the v) k φ)) r2)
      auxlist [ne..<ne+length auxlist])"
  unfolding wf_until_aux_def wf_until_auxlist_def qtable_mem_restr_UNIV ..

definition wf_matchF_aux :: "Formula.trace  _  nat  event_data list set 
      Formula.formula Regex.regex  event_data mlδaux  nat  nat  bool" where
  "wf_matchF_aux σ V n R I r aux ne k  (case to_mregex r of (mr, φs) 
      list_all2 (λx i. case x of (t, rels, rel)  t = τ σ i 
        list_all2 (λφ. qtable n (Formula.fv φ) (mem_restr R) (λv.
          Formula.sat σ V (map the v) i φ)) φs rels 
        qtable n (Formula.fv_regex r) (mem_restr R) (λv. (j. i  j  j < ne + length aux + k  mem (τ σ j - τ σ i) I 
          Regex.match (Formula.sat σ V (map the v)) r i j)) rel)
    aux [ne..<ne+length aux])"

definition wf_matchF_invar where
  "wf_matchF_invar σ V n R I r st i =
     (case st of (aux, Y)  aux  []  wf_matchF_aux σ V n R I r aux i 0 
     (case to_mregex r of (mr, φs)  ms  LPDs mr.
       qtable n (Formula.fv_regex r) (mem_restr R) (λv.
         Regex.match (Formula.sat σ V (map the v)) (from_mregex ms φs) i (i + length aux - 1)) (lookup Y ms)))"

definition lift_envs' :: "nat  event_data list set  event_data list set" where
  "lift_envs' b R = (λ(xs,ys). xs @ ys) ` ({xs. length xs = b} × R)"

fun formula_of_constraint :: "Formula.trm × bool × mconstraint × Formula.trm  Formula.formula" where
  "formula_of_constraint (t1, True, MEq, t2) = Formula.Eq t1 t2"
| "formula_of_constraint (t1, True, MLess, t2) = Formula.Less t1 t2"
| "formula_of_constraint (t1, True, MLessEq, t2) = Formula.LessEq t1 t2"
| "formula_of_constraint (t1, False, MEq, t2) = Formula.Neg (Formula.Eq t1 t2)"
| "formula_of_constraint (t1, False, MLess, t2) = Formula.Neg (Formula.Less t1 t2)"
| "formula_of_constraint (t1, False, MLessEq, t2) = Formula.Neg (Formula.LessEq t1 t2)"

inductive (in maux) wf_mformula :: "Formula.trace  nat  _  _ 
  nat  event_data list set  ('msaux, 'muaux) mformula  Formula.formula  bool"
  for σ j where
    Eq: "is_simple_eq t1 t2 
    xFormula.fv_trm t1. x < n  xFormula.fv_trm t2. x < n 
    wf_mformula σ j P V n R (MRel (eq_rel n t1 t2)) (Formula.Eq t1 t2)"
  | neq_Var: "x < n 
    wf_mformula σ j P V n R (MRel empty_table) (Formula.Neg (Formula.Eq (Formula.Var x) (Formula.Var x)))"
  | Pred: "xFormula.fv (Formula.Pred e ts). x < n 
    tset ts. Formula.is_Var t  Formula.is_Const t 
    wf_mformula σ j P V n R (MPred e ts) (Formula.Pred e ts)"
  | Let: "wf_mformula σ j P V m UNIV φ φ' 
    wf_mformula σ j (P(p  progress σ P φ' j))
      (V(p  λi. {v. length v = m  Formula.sat σ V v i φ'})) n R ψ ψ' 
    {0..<m}  Formula.fv φ'  b  m  m = Formula.nfv φ' 
    wf_mformula σ j P V n R (MLet p m φ ψ) (Formula.Let p φ' ψ')"
  | And: "wf_mformula σ j P V n R φ φ'  wf_mformula σ j P V n R ψ ψ' 
    if pos then χ = Formula.And φ' ψ'
      else χ = Formula.And φ' (Formula.Neg ψ')  Formula.fv ψ'  Formula.fv φ' 
    wf_mbuf2' σ P V j n R φ' ψ' buf 
    wf_mformula σ j P V n R (MAnd (fv φ') φ pos (fv ψ') ψ buf) χ"
  | AndAssign: "wf_mformula σ j P V n R φ φ' 
    x < n  x  Formula.fv φ'  Formula.fv_trm t  Formula.fv φ'  (x, t) = conf 
    ψ' = Formula.Eq (Formula.Var x) t  ψ' = Formula.Eq t (Formula.Var x) 
    wf_mformula σ j P V n R (MAndAssign φ conf) (Formula.And φ' ψ')"
  | AndRel: "wf_mformula σ j P V n R φ φ' 
    ψ' = formula_of_constraint conf 
    (let (t1, _, _, t2) = conf in Formula.fv_trm t1  Formula.fv_trm t2  Formula.fv φ') 
    wf_mformula σ j P V n R (MAndRel φ conf) (Formula.And φ' ψ')"
  | Ands: "list_all2 (λφ φ'. wf_mformula σ j P V n R φ φ') l (l_pos @ map remove_neg l_neg) 
    wf_mbufn (progress σ P (Formula.Ands l') j) (map (λψ. progress σ P ψ j) (l_pos @ map remove_neg l_neg)) (map (λψ i.
      qtable n (Formula.fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) i ψ)) (l_pos @ map remove_neg l_neg)) buf 
    (l_pos, l_neg) = partition safe_formula l' 
    l_pos  [] 
    list_all safe_formula (map remove_neg l_neg) 
    A_pos = map fv l_pos 
    A_neg = map fv l_neg 
    (set A_neg)  (set A_pos) 
    wf_mformula σ j P V n R (MAnds A_pos A_neg l buf) (Formula.Ands l')"
  | Or: "wf_mformula σ j P V n R φ φ'  wf_mformula σ j P V n R ψ ψ' 
    Formula.fv φ' = Formula.fv ψ' 
    wf_mbuf2' σ P V j n R φ' ψ' buf 
    wf_mformula σ j P V n R (MOr φ ψ buf) (Formula.Or φ' ψ')"
  | Neg: "wf_mformula σ j P V n R φ φ'  Formula.fv φ' = {} 
    wf_mformula σ j P V n R (MNeg φ) (Formula.Neg φ')"
  | Exists: "wf_mformula σ j P V (Suc n) (lift_envs R) φ φ' 
    wf_mformula σ j P V n R (MExists φ) (Formula.Exists φ')"
  | Agg: "wf_mformula σ j P V (b + n) (lift_envs' b R) φ φ' 
    y < n 
    y + b  Formula.fv φ' 
    {0..<b}  Formula.fv φ' 
    Formula.fv_trm f  Formula.fv φ' 
    g0 = (Formula.fv φ'  {0..<b}) 
    wf_mformula σ j P V n R (MAgg g0 y ω b f φ) (Formula.Agg y ω b f φ')"
  | Prev: "wf_mformula σ j P V n R φ φ' 
    first  j = 0 
    list_all2 (λi. qtable n (Formula.fv φ') (mem_restr R) (λv. Formula.sat σ V (map the v) i φ'))
      [min (progress σ P φ' j) (j-1)..<progress σ P φ' j] buf 
    list_all2 (λi t. t = τ σ i) [min (progress σ P φ' j) (j-1)..<j] nts 
    wf_mformula σ j P V n R (MPrev I φ first buf nts) (Formula.Prev I φ')"
  | Next: "wf_mformula σ j P V n R φ φ' 
    first  progress σ P φ' j = 0 
    list_all2 (λi t. t = τ σ i) [progress σ P φ' j - 1..<j] nts 
    wf_mformula σ j P V n R (MNext I φ first nts) (Formula.Next I φ')"
  | Since: "wf_mformula σ j P V n R φ φ'  wf_mformula σ j P V n R ψ ψ' 
    if args_pos args then φ'' = φ' else φ'' = Formula.Neg φ' 
    safe_formula φ'' = args_pos args 
    args_ivl args = I 
    args_n args = n 
    args_L args = Formula.fv φ' 
    args_R args = Formula.fv ψ' 
    Formula.fv φ'  Formula.fv ψ' 
    wf_mbuf2' σ P V j n R φ' ψ' buf 
    wf_ts σ P j φ' ψ' nts 
    wf_since_aux σ V R args φ' ψ' aux (progress σ P (Formula.Since φ'' I ψ') j) 
    wf_mformula σ j P V n R (MSince args φ ψ buf nts aux) (Formula.Since φ'' I ψ')"
  | Until: "wf_mformula σ j P V n R φ φ'  wf_mformula σ j P V n R ψ ψ' 
    if args_pos args then φ'' = φ' else φ'' = Formula.Neg φ' 
    safe_formula φ'' = args_pos args 
    args_ivl args = I 
    args_n args = n 
    args_L args = Formula.fv φ' 
    args_R args = Formula.fv ψ' 
    Formula.fv φ'  Formula.fv ψ' 
    wf_mbuf2' σ P V j n R φ' ψ' buf 
    wf_ts σ P j φ' ψ' nts 
    wf_until_aux σ V R args φ' ψ' aux (progress σ P (Formula.Until φ'' I ψ') j) 
    progress σ P (Formula.Until φ'' I ψ') j + length_muaux args aux = min (progress σ P φ' j) (progress σ P ψ' j) 
    wf_mformula σ j P V n R (MUntil args φ ψ buf nts aux) (Formula.Until φ'' I ψ')"
  | MatchP: "(case to_mregex r of (mr', φs') 
      list_all2 (wf_mformula σ j P V n R) φs φs'  mr = mr') 
    mrs = sorted_list_of_set (RPDs mr) 
    safe_regex Past Strict r 
    wf_mbufn' σ P V j n R r buf 
    wf_ts_regex σ P j r nts 
    wf_matchP_aux σ V n R I r aux (progress σ P (Formula.MatchP I r) j) 
    wf_mformula σ j P V n R (MMatchP I mr mrs φs buf nts aux) (Formula.MatchP I r)"
  | MatchF: "(case to_mregex r of (mr', φs') 
      list_all2 (wf_mformula σ j P V n R) φs φs'  mr = mr') 
    mrs = sorted_list_of_set (LPDs mr) 
    safe_regex Futu Strict r 
    wf_mbufn' σ P V j n R r buf 
    wf_ts_regex σ P j r nts 
    wf_matchF_aux σ V n R I r aux (progress σ P (Formula.MatchF I r) j) 0 
    progress σ P (Formula.MatchF I r) j + length aux = progress_regex σ P r j 
    wf_mformula σ j P V n R (MMatchF I mr mrs φs buf nts aux) (Formula.MatchF I r)"

definition (in maux) wf_mstate :: "Formula.formula  Formula.prefix  event_data list set  ('msaux, 'muaux) mstate  bool" where
  "wf_mstate φ π R st  mstate_n st = Formula.nfv φ  (σ. prefix_of π σ 
    mstate_i st = progress σ Map.empty φ (plen π) 
    wf_mformula σ (plen π) Map.empty Map.empty (mstate_n st) R (mstate_m st) φ)"


subsubsection ‹Initialisation›


lemma wf_mbuf2'_0: "pred_mapping (λx. x = 0) P  wf_mbuf2' σ P V 0 n R φ ψ ([], [])"
  unfolding wf_mbuf2'_def wf_mbuf2_def by simp

lemma wf_mbufn'_0: "to_mregex r = (mr, φs)  pred_mapping (λx. x = 0) P  wf_mbufn' σ P V 0 n R r (replicate (length φs) [])"
  unfolding wf_mbufn'_def wf_mbufn_def map_replicate_const[symmetric]
  by (auto simp: list_all3_map intro: list_all3_refl simp: Min_eq_iff progress_regex_def)

lemma wf_ts_0: "wf_ts σ P 0 φ ψ []"
  unfolding wf_ts_def by simp

lemma wf_ts_regex_0: "wf_ts_regex σ P 0 r []"
  unfolding wf_ts_regex_def by simp

lemma (in msaux) wf_since_aux_Nil: "Formula.fv φ'  Formula.fv ψ' 
  wf_since_aux σ V R (init_args I n (Formula.fv φ') (Formula.fv ψ') b) φ' ψ' (init_msaux (init_args I n (Formula.fv φ') (Formula.fv ψ') b)) 0"
  unfolding wf_since_aux_def by (auto intro!: valid_init_msaux)

lemma (in muaux) wf_until_aux_Nil: "Formula.fv φ'  Formula.fv ψ' 
  wf_until_aux σ V R (init_args I n (Formula.fv φ') (Formula.fv ψ') b) φ' ψ' (init_muaux (init_args I n (Formula.fv φ') (Formula.fv ψ') b)) 0"
  unfolding wf_until_aux_def wf_until_auxlist_def by (auto intro: valid_init_muaux)

lemma wf_matchP_aux_Nil: "wf_matchP_aux σ V n R I r [] 0"
  unfolding wf_matchP_aux_def by simp

lemma wf_matchF_aux_Nil: "wf_matchF_aux σ V n R I r [] 0 k"
  unfolding wf_matchF_aux_def by simp

lemma fv_regex_alt: "safe_regex m g r  Formula.fv_regex r = (φ  atms r. Formula.fv φ)"
  unfolding fv_regex_alt atms_def
  by (auto 0 3 dest: safe_regex_safe_formula)

lemmas to_mregex_atms =
  to_mregex_ok[THEN conjunct1, THEN equalityD1, THEN set_mp, rotated]

lemma (in maux) wf_minit0: "safe_formula φ  xFormula.fv φ. x < n 
  pred_mapping (λx. x = 0) P 
  wf_mformula σ 0 P V n R (minit0 n φ) φ"
proof (induction arbitrary: n R P V rule: safe_formula_induct)
  case (Eq_Const c d)
  then show ?case
    by (auto simp add: is_simple_eq_def simp del: eq_rel.simps intro!: wf_mformula.Eq)
next
  case (Eq_Var1 c x)
  then show ?case
    by (auto simp add: is_simple_eq_def simp del: eq_rel.simps intro!: wf_mformula.Eq)
next
  case (Eq_Var2 c x)
  then show ?case
    by (auto simp add: is_simple_eq_def simp del: eq_rel.simps intro!: wf_mformula.Eq)
next
  case (neq_Var x y)
  then show ?case by (auto intro!: wf_mformula.neq_Var)
next
  case (Pred e ts)
  then show ?case by (auto intro!: wf_mformula.Pred)
next
  case (Let p φ ψ)
  with fvi_less_nfv show ?case
    by (auto simp: pred_mapping_alt dom_def intro!: wf_mformula.Let Let(4,5))
next
  case (And_assign φ ψ)
  then have 1: "xfv ψ. x < n" by simp
  from 1 safe_assignment (fv φ) ψ
  obtain x t where
    "x < n" "x  fv φ" "fv_trm t  fv φ"
    "ψ = Formula.Eq (Formula.Var x) t  ψ = Formula.Eq t (Formula.Var x)"
    unfolding safe_assignment_def by (force split: formula.splits trm.splits)
  with And_assign show ?case
    by (auto intro!: wf_mformula.AndAssign split: trm.splits)
next
  case (And_safe φ ψ)
  then show ?case by (auto intro!: wf_mformula.And wf_mbuf2'_0)
next
  case (And_constraint φ ψ)
  from fv ψ  fv φ is_constraint ψ
  obtain t1 p c t2 where
    "(t1, p, c, t2) = split_constraint ψ"
    "formula_of_constraint (split_constraint ψ) = ψ"
    "fv_trm t1  fv_trm t2  fv φ"
    by (induction rule: is_constraint.induct) auto
  with And_constraint show ?case
    by (auto 0 3 intro!: wf_mformula.AndRel)
next
  case (And_Not φ ψ)
  then show ?case by (auto intro!: wf_mformula.And wf_mbuf2'_0)
next
  case (Ands l pos neg)
  note posneg = "Ands.hyps"(1)
  let ?wf_minit = "λx. wf_mformula σ 0 P V n R (minit0 n x)"
  let ?pos = "filter safe_formula l"
  let ?neg = "filter (Not  safe_formula) l"
  have "list_all2 ?wf_minit ?pos pos"
    using Ands.IH(1) Ands.prems posneg by (auto simp: list_all_iff intro!: list.rel_refl_strong)
  moreover have "list_all2 ?wf_minit (map remove_neg ?neg) (map remove_neg neg)"
    using Ands.IH(2) Ands.prems posneg by (auto simp: list.rel_map list_all_iff intro!: list.rel_refl_strong)
  moreover have "list_all3 (λ_ _ _. True) (?pos @ map remove_neg ?neg) (?pos @ map remove_neg ?neg) l"
    by (auto simp: list_all3_conv_all_nth comp_def sum_length_filter_compl)
  moreover have "l  []  (MIN φset l. (0 :: nat)) = 0"
    by (cases l) (auto simp: Min_eq_iff)
  ultimately show ?case using Ands.hyps Ands.prems(2)
    by (auto simp: wf_mbufn_def list_all3_map list.rel_map map_replicate_const[symmetric] subset_eq
        map_map[symmetric] map_append[symmetric] simp del: map_map map_append
        intro!: wf_mformula.Ands list_all2_appendI)
next
  case (Neg φ)
  then show ?case by (auto intro!: wf_mformula.Neg)
next
  case (Or φ ψ)
  then show ?case by (auto intro!: wf_mformula.Or wf_mbuf2'_0)
next
  case (Exists φ)
  then show ?case by (auto simp: fvi_Suc_bound intro!: wf_mformula.Exists)
next
  case (Agg y ω b f φ)
  then show ?case by (auto intro!: wf_mformula.Agg Agg.IH fvi_plus_bound)
next
  case (Prev I φ)
  thm wf_mformula.Prev[where P=P]
  then show ?case by (auto intro!: wf_mformula.Prev)
next
  case (Next I φ)
  then show ?case by (auto intro!: wf_mformula.Next)
next
  case (Since φ I ψ)
  then show ?case
    using wf_since_aux_Nil
    by (auto simp add: init_args_def intro!: wf_mformula.Since wf_mbuf2'_0 wf_ts_0)
next
  case (Not_Since φ I ψ)
  then show ?case
    using wf_since_aux_Nil
    by (auto simp add: init_args_def intro!: wf_mformula.Since wf_mbuf2'_0 wf_ts_0)
next
  case (Until φ I ψ)
  then show ?case
    using valid_length_muaux[OF valid_init_muaux[OF Until(1)]] wf_until_aux_Nil
    by (auto simp add: init_args_def simp del: progress_simps intro!: wf_mformula.Until wf_mbuf2'_0 wf_ts_0)
next
  case (Not_Until φ I ψ)
  then show ?case
    using valid_length_muaux[OF valid_init_muaux[OF Not_Until(1)]] wf_until_aux_Nil
    by (auto simp add: init_args_def simp del: progress_simps intro!: wf_mformula.Until wf_mbuf2'_0 wf_ts_0)
next
  case (MatchP I r)
  then show ?case
    by (auto simp: list.rel_map fv_regex_alt simp del: progress_simps split: prod.split
        intro!: wf_mformula.MatchP list.rel_refl_strong wf_mbufn'_0 wf_ts_regex_0 wf_matchP_aux_Nil
        dest!: to_mregex_atms)
next
  case (MatchF I r)
  then show ?case
    by (auto simp: list.rel_map fv_regex_alt progress_le Min_eq_iff progress_regex_def
        simp del: progress_simps split: prod.split
        intro!: wf_mformula.MatchF list.rel_refl_strong wf_mbufn'_0 wf_ts_regex_0 wf_matchF_aux_Nil
        dest!: to_mregex_atms)
qed

lemma (in maux) wf_mstate_minit: "safe_formula φ  wf_mstate φ pnil R (minit φ)"
  unfolding wf_mstate_def minit_def Let_def
  by (auto intro!: wf_minit0 fvi_less_nfv)


subsubsection ‹Evaluation›

lemma match_wf_tuple: "Some f = match ts xs 
  wf_tuple n (tset ts. Formula.fv_trm t) (Table.tabulate f 0 n)"
  by (induction ts xs arbitrary: f rule: match.induct)
    (fastforce simp: wf_tuple_def split: if_splits option.splits)+

lemma match_fvi_trm_None: "Some f = match ts xs  tset ts. x  Formula.fv_trm t  f x = None"
  by (induction ts xs arbitrary: f rule: match.induct) (auto split: if_splits option.splits)

lemma match_fvi_trm_Some: "Some f = match ts xs  t  set ts  x  Formula.fv_trm t  f x  None"
  by (induction ts xs arbitrary: f rule: match.induct) (auto split: if_splits option.splits)

lemma match_eval_trm: "tset ts. iFormula.fv_trm t. i < n  Some f = match ts xs 
    map (Formula.eval_trm (Table.tabulate (λi. the (f i)) 0 n)) ts = xs"
proof (induction ts xs arbitrary: f rule: match.induct)
  case (3 x ts y ys)
  from 3(1)[symmetric] 3(2,3) show ?case
    by (auto 0 3 dest: match_fvi_trm_Some sym split: option.splits if_splits intro!: eval_trm_fv_cong)
qed (auto split: if_splits)

lemma wf_tuple_tabulate_Some: "wf_tuple n A (Table.tabulate f 0 n)  x  A  x < n  y. f x = Some y"
  unfolding wf_tuple_def by auto

lemma ex_match: "wf_tuple n (tset ts. Formula.fv_trm t) v 
  tset ts. (xFormula.fv_trm t. x < n)  (Formula.is_Var t  Formula.is_Const t) 
  f. match ts (map (Formula.eval_trm (map the v)) ts) = Some f  v = Table.tabulate f 0 n"
proof (induction ts "map (Formula.eval_trm (map the v)) ts" arbitrary: v rule: match.induct)
  case (3 x ts y ys)
  then show ?case
  proof (cases "x  (tset ts. Formula.fv_trm t)")
    case True
    with 3 show ?thesis
      by (auto simp: insert_absorb dest!: wf_tuple_tabulate_Some meta_spec[of _ v])
  next
    case False
    with 3(3,4) have
      *: "map (Formula.eval_trm (map the v)) ts = map (Formula.eval_trm (map the (v[x := None]))) ts"
      by (auto simp: wf_tuple_def nth_list_update intro!: eval_trm_fv_cong)
    from False 3(2-4) obtain f where
      "match ts (map (Formula.eval_trm (map the v)) ts) = Some f" "v[x := None] = Table.tabulate f 0 n"
      unfolding *
      by (atomize_elim, intro 3(1)[of "v[x := None]"])
        (auto simp: wf_tuple_def nth_list_update intro!: eval_trm_fv_cong)
    moreover from False this have "f x = None" "length v = n"
      by (auto dest: match_fvi_trm_None[OF sym] arg_cong[of _ _ length])
    ultimately show ?thesis using 3(3)
      by (auto simp: list_eq_iff_nth_eq wf_tuple_def)
  qed
qed (auto simp: wf_tuple_def intro: nth_equalityI)

lemma eq_rel_eval_trm: "v  eq_rel n t1 t2  is_simple_eq t1 t2 
  xFormula.fv_trm t1  Formula.fv_trm t2. x < n 
  Formula.eval_trm (map the v) t1 = Formula.eval_trm (map the v) t2"
  by (cases t1; cases t2) (simp_all add: is_simple_eq_def singleton_table_def split: if_splits)

lemma in_eq_rel: "wf_tuple n (Formula.fv_trm t1  Formula.fv_trm t2) v 
  is_simple_eq t1 t2 
  Formula.eval_trm (map the v) t1 = Formula.eval_trm (map the v) t2 
  v  eq_rel n t1 t2"
  by (cases t1; cases t2)
    (auto simp: is_simple_eq_def singleton_table_def wf_tuple_def unit_table_def
      intro!: nth_equalityI split: if_splits)

lemma table_eq_rel: "is_simple_eq t1 t2 
  table n (Formula.fv_trm t1  Formula.fv_trm t2) (eq_rel n t1 t2)"
  by (cases t1; cases t2; simp add: is_simple_eq_def)

lemma wf_tuple_Suc_fviD: "wf_tuple (Suc n) (Formula.fvi b φ) v  wf_tuple n (Formula.fvi (Suc b) φ) (tl v)"
  unfolding wf_tuple_def by (simp add: fvi_Suc nth_tl)

lemma table_fvi_tl: "table (Suc n) (Formula.fvi b φ) X  table n (Formula.fvi (Suc b) φ) (tl ` X)"
  unfolding table_def by (auto intro: wf_tuple_Suc_fviD)

lemma wf_tuple_Suc_fvi_SomeI: "0  Formula.fvi b φ  wf_tuple n (Formula.fvi (Suc b) φ) v 
  wf_tuple (Suc n) (Formula.fvi b φ) (Some x # v)"
  unfolding wf_tuple_def
  by (auto simp: fvi_Suc less_Suc_eq_0_disj)

lemma wf_tuple_Suc_fvi_NoneI: "0  Formula.fvi b φ  wf_tuple n (Formula.fvi (Suc b) φ) v 
  wf_tuple (Suc n) (Formula.fvi b φ) (None # v)"
  unfolding wf_tuple_def
  by (auto simp: fvi_Suc less_Suc_eq_0_disj)

lemma qtable_project_fv: "qtable (Suc n) (fv φ) (mem_restr (lift_envs R)) P X 
    qtable n (Formula.fvi (Suc 0) φ) (mem_restr R)
      (λv. x. P ((if 0  fv φ then Some x else None) # v)) (tl ` X)"
  using neq0_conv by (fastforce simp: image_iff Bex_def fvi_Suc elim!: qtable_cong dest!: qtable_project)

lemma mem_restr_lift_envs'_append[simp]:
  "length xs = b  mem_restr (lift_envs' b R) (xs @ ys) = mem_restr R ys"
  unfolding mem_restr_def lift_envs'_def
  by (auto simp: list_all2_append list.rel_map intro!: exI[where x="map the xs"] list.rel_refl)

lemma nth_list_update_alt: "xs[i := x] ! j = (if i < length xs  i = j then x else xs ! j)"
  by auto

lemma wf_tuple_upd_None: "wf_tuple n A xs  A - {i} = B  wf_tuple n B (xs[i:=None])"
  unfolding wf_tuple_def
  by (auto simp: nth_list_update_alt)

lemma mem_restr_upd_None: "mem_restr R xs  mem_restr R (xs[i:=None])"
  unfolding mem_restr_def
  by (auto simp: list_all2_conv_all_nth nth_list_update_alt)

lemma mem_restr_dropI: "mem_restr (lift_envs' b R) xs  mem_restr R (drop b xs)"
  unfolding mem_restr_def lift_envs'_def
  by (auto simp: append_eq_conv_conj list_all2_append2)

lemma mem_restr_dropD:
  assumes "b  length xs" and "mem_restr R (drop b xs)"
  shows "mem_restr (lift_envs' b R) xs"
proof -
  let ?R = "λa b. a  None  a = Some b"
  from assms(2) obtain v where "v  R" and "list_all2 ?R (drop b xs) v"
    unfolding mem_restr_def ..
  show ?thesis unfolding mem_restr_def proof
    have "list_all2 ?R (take b xs) (map the (take b xs))"
      by (auto simp: list.rel_map intro!: list.rel_refl)
    moreover note list_all2 ?R (drop b xs) v
    ultimately have "list_all2 ?R (take b xs @ drop b xs) (map the (take b xs) @ v)"
      by (rule list_all2_appendI)
    then show "list_all2 ?R xs (map the (take b xs) @ v)" by simp
    show "map the (take b xs) @ v  lift_envs' b R"
      unfolding lift_envs'_def using assms(1) v  R by auto
  qed
qed

lemma wf_tuple_append: "wf_tuple a {x  A. x < a} xs 
  wf_tuple b {x - a | x. x  A  x  a} ys 
  wf_tuple (a + b) A (xs @ ys)"
  unfolding wf_tuple_def by (auto simp: nth_append eq_diff_iff)

lemma wf_tuple_map_Some: "length xs = n  {0..<n}  A  wf_tuple n A (map Some xs)"
  unfolding wf_tuple_def by auto

lemma wf_tuple_drop: "wf_tuple (b + n) A xs  {x - b | x. x  A  x  b} = B 
  wf_tuple n B (drop b xs)"
  unfolding wf_tuple_def by force

lemma ecard_image: "inj_on f A  ecard (f ` A) = ecard A"
  unfolding ecard_def by (auto simp: card_image dest: finite_imageD)

lemma meval_trm_eval_trm: "wf_tuple n A x  fv_trm t  A  iA. i < n 
    meval_trm t x = Formula.eval_trm (map the x) t"
  unfolding wf_tuple_def
  by (induction t) simp_all

lemma list_update_id: "xs ! i = z  xs[i:=z] = xs"
  by (induction xs arbitrary: i) (auto split: nat.split)

lemma qtable_wf_tupleD: "qtable n A P Q X  xX. wf_tuple n A x"
  unfolding qtable_def table_def by blast

lemma qtable_eval_agg:
  assumes inner: "qtable (b + n) (Formula.fv φ) (mem_restr (lift_envs' b R))
      (λv. Formula.sat σ V (map the v) i φ) rel"
    and n: "xFormula.fv (Formula.Agg y ω b f φ). x < n"
    and fresh: "y + b  Formula.fv φ"
    and b_fv: "{0..<b}  Formula.fv φ"
    and f_fv: "Formula.fv_trm f  Formula.fv φ"
    and g0: "g0 = (Formula.fv φ  {0..<b})"
  shows "qtable n (Formula.fv (Formula.Agg y ω b f φ)) (mem_restr R)
      (λv. Formula.sat σ V (map the v) i (Formula.Agg y ω b f φ)) (eval_agg n g0 y ω b f rel)"
    (is "qtable _ ?fv _ ?Q ?rel'")
proof -
  define M where "M = (λv. {(x, ecard Zs) | x Zs.
      Zs = {zs. length zs = b  Formula.sat σ V (zs @ v) i φ  Formula.eval_trm (zs @ v) f = x} 
      Zs  {}})"
  have f_fvi: "Formula.fvi_trm b f  Formula.fvi b φ"
    using f_fv by (auto simp: fvi_trm_iff_fv_trm[where b=b] fvi_iff_fv[where b=b])
  show ?thesis proof (cases "g0  rel = empty_table")
    case True
    then have [simp]: "Formula.fvi b φ = {}"
      by (auto simp: g0 fvi_iff_fv(1)[where b=b])
    then have [simp]: "Formula.fvi_trm b f = {}"
      using f_fvi by auto
    show ?thesis proof (rule qtableI)
      show "table n ?fv ?rel'" by (simp add: eval_agg_def True)
    next
      fix v
      assume "wf_tuple n ?fv v" "mem_restr R v"
      have "¬ Formula.sat σ V (zs @ map the v) i φ" if [simp]: "length zs = b" for zs
      proof -
        let ?zs = "map2 (λz i. if i  Formula.fv φ then Some z else None) zs [0..<b]"
        have "wf_tuple b {x  fv φ. x < b} ?zs"
          by (simp add: wf_tuple_def)
        then have "wf_tuple (b + n) (Formula.fv φ) (?zs @ v[y:=None])"
          using wf_tuple n ?fv v True
          by (auto simp: g0 intro!: wf_tuple_append wf_tuple_upd_None)
        then have "¬ Formula.sat σ V (map the (?zs @ v[y:=None])) i φ"
          using True mem_restr R v
          by (auto simp del: map_append dest!: in_qtableI[OF inner, rotated -1]
              intro!: mem_restr_upd_None)
        also have "Formula.sat σ V (map the (?zs @ v[y:=None])) i φ  Formula.sat σ V (zs @ map the v) i φ"
          using True by (auto simp: g0 nth_append intro!: sat_fv_cong)
        finally show ?thesis .
      qed
      then have M_empty: "M (map the v) = {}"
        unfolding M_def by blast
      show "Formula.sat σ V (map the v) i (Formula.Agg y ω b f φ)"
        if "v  eval_agg n g0 y ω b f rel"
        using M_empty True that n
        by (simp add: M_def eval_agg_def g0 singleton_table_def)
      have "v  singleton_table n y (the (v ! y))" "length v = n"
        using wf_tuple n ?fv v unfolding wf_tuple_def singleton_table_def
        by (auto simp add: tabulate_alt map_nth
            intro!: trans[OF map_cong[where g="(!) v", simplified nth_map, OF refl], symmetric])
      then show "v  eval_agg n g0 y ω b f rel"
        if "Formula.sat σ V (map the v) i (Formula.Agg y ω b f φ)"
        using M_empty True that n
        by (simp add: M_def eval_agg_def g0)
    qed
  next
    case non_default_case: False
    have union_fv: "{0..<b}  (λx. x + b) ` Formula.fvi b φ = fv φ"
      using b_fv
      by (auto simp: fvi_iff_fv(1)[where b=b] intro!: image_eqI[where b=x and x="x - b" for x])
    have b_n: "xfv φ. x < b + n"
    proof
      fix x assume "x  fv φ"
      show "x < b + n" proof (cases "x  b")
        case True
        with x  fv φ have "x - b  ?fv"
          by (simp add: fvi_iff_fv(1)[where b=b])
        then show ?thesis using n f_fvi by (auto simp: Un_absorb2)
      qed simp
    qed

    define M' where "M' = (λk. let group = Set.filter (λx. drop b x = k) rel;
        images = meval_trm f ` group
      in (λy. (y, ecard (Set.filter (λx. meval_trm f x = y) group))) ` images)"
    have M'_M: "M' (drop b x) = M (map the (drop b x))" if "x  rel" "mem_restr (lift_envs' b R) x" for x
    proof -
      from that have wf_x: "wf_tuple (b + n) (fv φ) x"
        by (auto elim!: in_qtableE[OF inner])
      then have wf_zs_x: "wf_tuple (b + n) (fv φ) (map Some zs @ drop b x)"
        if "length zs = b" for zs
        using that b_fv
        by (auto intro!: wf_tuple_append wf_tuple_map_Some wf_tuple_drop)
      have 1: "(length zs = b  Formula.sat σ V (zs @ map the (drop b x)) i φ 
          Formula.eval_trm (zs @ map the (drop b x)) f = y) 
        (a. a  rel  take b a = map Some zs  drop b a = drop b x  meval_trm f a = y)"
        (is "?A  (a. ?B a)") for y zs
      proof (intro iffI conjI)
        assume ?A
        then have "?B (map Some zs @ drop (length zs) x)"
          using in_qtableI[OF inner wf_zs_x] mem_restr (lift_envs' b R) x
            meval_trm_eval_trm[OF wf_zs_x f_fv b_n]
          by (auto intro!: mem_restr_dropI)
        then show "a. ?B a" ..
      next
        assume "a. ?B a"
        then obtain a where "?B a" ..
        then have "a  rel" and a_eq: "a = map Some zs @ drop b x"
          using append_take_drop_id[of b a] by auto
        then have "length a = b + n"
          using inner unfolding qtable_def table_def
          by (blast intro!: wf_tuple_length)
        then show "length zs = b"
          using wf_tuple_length[OF wf_x] unfolding a_eq by simp
        then have "mem_restr (lift_envs' b R) a"
          using mem_restr _ x unfolding a_eq by (auto intro!: mem_restr_dropI)
        then show "Formula.sat σ V (zs @ map the (drop b x)) i φ"
          using in_qtableE[OF inner a  rel]
          by (auto simp: a_eq sat_fv_cong[THEN iffD1, rotated -1])
        from ?B a show "Formula.eval_trm (zs @ map the (drop b x)) f = y"
          using meval_trm_eval_trm[OF wf_zs_x f_fv b_n, OF length zs = b]
          unfolding a_eq by simp
      qed
      have 2: "map Some (map the (take b a)) = take b a" if "a  rel" for a
        using that b_fv inner[THEN qtable_wf_tupleD]
        unfolding table_def wf_tuple_def
        by (auto simp: list_eq_iff_nth_eq)
      have 3: "ecard {zs. a. a  rel  take b a = map Some zs  drop b a = drop b x  P a} =
        ecard {a. a  rel  drop b a = drop b x  P a}" (is "ecard ?A = ecard ?B") for P
      proof -
        have "ecard ?A = ecard ((λzs. map Some zs @ drop b x) ` ?A)"
          by (auto intro!: ecard_image[symmetric] inj_onI)
        also have "(λzs. map Some zs @ drop b x) ` ?A = ?B"
          by (subst (1 2) eq_commute) (auto simp: image_iff, metis "2" append_take_drop_id)
        finally show ?thesis .
      qed
      show ?thesis
        unfolding M_def M'_def
        by (auto simp: non_default_case Let_def image_def Set.filter_def 1 3, metis "2")
    qed
    have drop_lift: "mem_restr (lift_envs' b R) x" if "x  rel" "mem_restr R ((drop b x)[y:=z])" for x z
    proof -
      have "(drop b x)[y:=None] = (drop b x)[y:=drop b x ! y]" proof -
        from x  rel have "drop b x ! y = None"
          using fresh n inner[THEN qtable_wf_tupleD]
          by (simp add: add.commute wf_tuple_def)
        then show ?thesis by simp
      qed
      then have "(drop b x)[y:=None] = drop b x" by simp
      moreover from x  rel have "length x = b + n"
        using inner[THEN qtable_wf_tupleD]
        by (simp add: wf_tuple_def)
      moreover from that(2) have "mem_restr R ((drop b x)[y:=z, y:=None])"
        by (rule mem_restr_upd_None)
      ultimately show ?thesis
        by (auto intro!: mem_restr_dropD)
    qed

    {
      fix v
      assume "mem_restr R v"
      have "v  (λk. k[y:=Some (eval_agg_op ω (M' k))]) ` drop b ` rel 
          v  (λk. k[y:=Some (eval_agg_op ω (M (map the k)))]) ` drop b ` rel"
        (is "v  ?A  v  ?B")
      proof
        assume "v  ?A"
        then obtain v' where *: "v'  rel" "v = (drop b v')[y:=Some (eval_agg_op ω (M' (drop b v')))]"
          by auto
        then have "M' (drop b v') = M (map the (drop b v'))"
          using mem_restr R v by (auto intro!: M'_M drop_lift)
        with * show "v  ?B" by simp
      next
        assume "v  ?B"
        then obtain v' where *: "v'  rel" "v = (drop b v')[y:=Some (eval_agg_op ω (M (map the (drop b v'))))]"
          by auto
        then have "M (map the (drop b v')) = M' (drop b v')"
          using mem_restr R v by (auto intro!: M'_M[symmetric] drop_lift)
        with * show "v  ?A" by simp
      qed
      then have "v  eval_agg n g0 y ω b f rel  v  (λk. k[y:=Some (eval_agg_op ω (M (map the k)))]) ` drop b ` rel"
        by (simp add: non_default_case eval_agg_def M'_def Let_def)
    }
    note alt = this

    show ?thesis proof (rule qtableI)
      show "table n ?fv ?rel'"
        using inner[THEN qtable_wf_tupleD] n f_fvi
        by (auto simp: eval_agg_def non_default_case table_def wf_tuple_def Let_def nth_list_update
            fvi_iff_fv[where b=b] add.commute)
    next
      fix v
      assume "wf_tuple n ?fv v" "mem_restr R v"
      then have length_v: "length v = n" by (simp add: wf_tuple_def)

      show "Formula.sat σ V (map the v) i (Formula.Agg y ω b f φ)"
        if "v  eval_agg n g0 y ω b f rel"
      proof -
        from that obtain v' where "v'  rel"
          "v = (drop b v')[y:=Some (eval_agg_op ω (M (map the (drop b v'))))]"
          using alt[OF mem_restr R v] by blast
        then have length_v': "length v' = b + n"
          using inner[THEN qtable_wf_tupleD]
          by (simp add: wf_tuple_def)
        have "Formula.sat σ V (map the v') i φ"
          using v'  rel mem_restr R v
          by (auto simp: v = _ elim!: in_qtableE[OF inner] intro!: drop_lift v'  rel)
        then have "Formula.sat σ V (map the (take b v') @ map the v) i φ"
        proof (rule sat_fv_cong[THEN iffD1, rotated], intro ballI)
          fix x
          assume "x  fv φ"
          then have "x  y + b" using fresh by blast
          moreover have "x < length v'"
            using x  fv φ b_n by (simp add: length_v')
          ultimately show "map the v' ! x = (map the (take b v') @ map the v) ! x"
            by (auto simp: v = _ nth_append)
        qed
        then have 1: "M (map the v)  {}" by (force simp: M_def length_v')

        have "y < length (drop b v')" using n by (simp add: length_v')
        moreover have "Formula.sat σ V (zs @ map the v) i φ 
          Formula.sat σ V (zs @ map the (drop b v')) i φ" if "length zs = b" for zs
        proof (intro sat_fv_cong ballI)
          fix x
          assume "x  fv φ"
          then have "x  y + b" using fresh by blast
          moreover have "x < length v'"
            using x  fv φ b_n by (simp add: length_v')
          ultimately show "(zs @ map the v) ! x = (zs @ map the (drop b v')) ! x"
            by (auto simp: v = _ that nth_append)
        qed
        moreover have "Formula.eval_trm (zs @ map the v) f =
          Formula.eval_trm (zs @ map the (drop b v')) f" if "length zs = b" for zs
        proof (intro eval_trm_fv_cong ballI)
          fix x
          assume "x  fv_trm f"
          then have "x  y + b" using f_fv fresh by blast
          moreover have "x < length v'"
            using x  fv_trm f f_fv b_n by (auto simp: length_v')
          ultimately show "(zs @ map the v) ! x = (zs @ map the (drop b v')) ! x"
            by (auto simp: v = _ that nth_append)
        qed
        ultimately have "map the v ! y = eval_agg_op ω (M (map the v))"
          by (simp add: M_def v = _ conj_commute cong: conj_cong)
        with 1 show ?thesis by (auto simp: M_def)
      qed

      show "v  eval_agg n g0 y ω b f rel"
        if sat_Agg: "Formula.sat σ V (map the v) i (Formula.Agg y ω b f φ)"
      proof -
        obtain zs where "length zs = b" and "map Some zs @ v[y:=None]  rel"
        proof (cases "fv φ  {0..<b}")
          case True
          with non_default_case have "rel  empty_table" by (simp add: g0)
          then obtain x where "x  rel" by auto
          have "(i < n. (v[y:=None]) ! i = None)"
            using True wf_tuple n ?fv v f_fv
            by (fastforce simp: wf_tuple_def fvi_iff_fv[where b=b] fvi_trm_iff_fv_trm[where b=b])
          moreover have x: "(i < n. drop b x ! i = None)  length x = b + n"
            using True x  rel inner[THEN qtable_wf_tupleD] f_fv
            by (auto simp: wf_tuple_def)
          ultimately have "v[y:=None] = drop b x"
            unfolding list_eq_iff_nth_eq by (auto simp: length_v)
          with x  rel have "take b x @ v[y:=None]  rel" by simp
          moreover have "map (Some  the) (take b x) = take b x"
            using True x  rel inner[THEN qtable_wf_tupleD] b_fv
            by (subst map_cong[where g=id, OF refl]) (auto simp: wf_tuple_def in_set_conv_nth)
          ultimately have "map Some (map the (take b x)) @ v[y:=None]  rel" by simp
          then show thesis using x[THEN conjunct2] by (fastforce intro!: that[rotated])
        next
          case False
          with sat_Agg obtain zs where "length zs = b" and "Formula.sat σ V (zs @ map the v) i φ"
            by auto
          then have "Formula.sat σ V (zs @ map the (v[y:=None])) i φ"
            using fresh
            by (auto simp: map_update not_less nth_append elim!: sat_fv_cong[THEN iffD1, rotated]
                intro!: nth_list_update_neq[symmetric])
          then have "map Some zs @ v[y:=None]  rel"
            using b_fv f_fv fresh
            by (auto intro!: in_qtableI[OF inner] wf_tuple_append wf_tuple_map_Some
                wf_tuple_upd_None wf_tuple n ?fv v mem_restr_upd_None mem_restr R v
                simp: length zs = b set_eq_iff fvi_iff_fv[where b=b] fvi_trm_iff_fv_trm[where b=b])
              force+
          with that length zs = b show thesis by blast
        qed
        then have 1: "v[y:=None]  drop b ` rel" by (intro image_eqI) auto

        have y_length: "y < length v" using n by (simp add: length_v)
        moreover have "Formula.sat σ V (zs @ map the (v[y:=None])) i φ 
          Formula.sat σ V (zs @ map the v) i φ" if "length zs = b" for zs
        proof (intro sat_fv_cong ballI)
          fix x
          assume "x  fv φ"
          then have "x  y + b" using fresh by blast
          moreover have "x < b + length v"
            using x  fv φ b_n by (simp add: length_v)
          ultimately show "(zs @ map the (v[y:=None])) ! x = (zs @ map the v) ! x"
            by (auto simp: that nth_append)
        qed
        moreover have "Formula.eval_trm (zs @ map the (v[y:=None])) f =
          Formula.eval_trm (zs @ map the v) f" if "length zs = b" for zs
        proof (intro eval_trm_fv_cong ballI)
          fix x
          assume "x  fv_trm f"
          then have "x  y + b" using f_fv fresh by blast
          moreover have "x < b + length v"
            using x  fv_trm f f_fv b_n by (auto simp: length_v)
          ultimately show "(zs @ map the (v[y:=None])) ! x = (zs @ map the v) ! x"
            by (auto simp: that nth_append)
        qed
        ultimately have "map the v ! y = eval_agg_op ω (M (map the (v[y:=None])))"
          using sat_Agg by (simp add: M_def cong: conj_cong) (simp cong: rev_conj_cong)
        then have 2: "v ! y = Some (eval_agg_op ω (M (map the (v[y:=None]))))"
          using wf_tuple n ?fv v y_length by (auto simp add: wf_tuple_def)
        show ?thesis
          unfolding alt[OF mem_restr R v]
          by (rule image_eqI[where x="v[y:=None]"]) (use 1 2 in auto simp: y_length list_update_id)
      qed
    qed
  qed
qed

lemma mprev: "mprev_next I xs ts = (ys, xs', ts') 
  list_all2 P [i..<j'] xs  list_all2 (λi t. t = τ σ i) [i..<j] ts  i  j'  i < j 
  list_all2 (λi X. if mem (τ σ (Suc i) - τ σ i) I then P i X else X = empty_table)
    [i..<min j' (j-1)] ys 
  list_all2 P [min j' (j-1)..<j'] xs' 
  list_all2 (λi t. t = τ σ i) [min j' (j-1)..<j] ts'"
proof (induction I xs ts arbitrary: i ys xs' ts' rule: mprev_next.induct)
  case (1 I ts)
  then have "min j' (j-1) = i" by auto
  with 1 show ?case by auto
next
  case (3 I v v' t)
  then have "min j' (j-1) = i" by (auto simp: list_all2_Cons2 upt_eq_Cons_conv)
  with 3 show ?case by auto
next
  case (4 I x xs t t' ts)
  from 4(1)[of "tl ys" xs' ts' "Suc i"] 4(2-6) show ?case
    by (auto simp add: list_all2_Cons2 upt_eq_Cons_conv Suc_less_eq2
        elim!: list.rel_mono_strong split: prod.splits if_splits)
qed simp

lemma mnext: "mprev_next I xs ts = (ys, xs', ts') 
  list_all2 P [Suc i..<j'] xs  list_all2 (λi t. t = τ σ i) [i..<j] ts  Suc i  j'  i < j 
  list_all2 (λi X. if mem (τ σ (Suc i) - τ σ i) I then P (Suc i) X else X = empty_table)
    [i..<min (j'-1) (j-1)] ys 
  list_all2 P [Suc (min (j'-1) (j-1))..<j'] xs' 
  list_all2 (λi t. t = τ σ i) [min (j'-1) (j-1)..<j] ts'"
proof (induction I xs ts arbitrary: i ys xs' ts' rule: mprev_next.induct)
  case (1 I ts)
  then have "min (j' - 1) (j-1) = i" by auto
  with 1 show ?case by auto
next
  case (3 I v v' t)
  then have "min (j' - 1) (j-1) = i" by (auto simp: list_all2_Cons2 upt_eq_Cons_conv)
  with 3 show ?case by auto
next
  case (4 I x xs t t' ts)
  from 4(1)[of "tl ys" xs' ts' "Suc i"] 4(2-6) show ?case
    by (auto simp add: list_all2_Cons2 upt_eq_Cons_conv Suc_less_eq2
        elim!: list.rel_mono_strong split: prod.splits if_splits) (* slow 10 sec *)
qed simp

lemma in_foldr_UnI: "x  A  A  set xs  x  foldr (∪) xs {}"
  by (induction xs) auto

lemma in_foldr_UnE: "x  foldr (∪) xs {}  (A. A  set xs  x  A  P)  P"
  by (induction xs) auto

lemma sat_the_restrict: "fv φ  A  Formula.sat σ V (map the (restrict A v)) i φ = Formula.sat σ V (map the v) i φ"
  by (rule sat_fv_cong) (auto intro!: map_the_restrict)

lemma eps_the_restrict: "fv_regex r  A  Regex.eps (Formula.sat σ V (map the (restrict A v))) i r = Regex.eps (Formula.sat σ V (map the v)) i r"
  by (rule eps_fv_cong) (auto intro!: map_the_restrict)

lemma sorted_wrt_filter[simp]: "sorted_wrt R xs  sorted_wrt R (filter P xs)"
  by (induct xs) auto

lemma concat_map_filter[simp]:
  "concat (map f (filter P xs)) = concat (map (λx. if P x then f x else []) xs)"
  by (induct xs) auto

lemma map_filter_alt:
  "map f (filter P xs) = concat (map (λx. if P x then [f x] else []) xs)"
  by (induct xs) auto

lemma (in maux) update_since:
  assumes pre: "wf_since_aux σ V R args φ ψ aux ne"
    and qtable1: "qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) ne φ) rel1"
    and qtable2: "qtable n (Formula.fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) ne ψ) rel2"
    and result_eq: "(rel, aux') = update_since args rel1 rel2 (τ σ ne) aux"
    and fvi_subset: "Formula.fv φ  Formula.fv ψ"
    and args_ivl: "args_ivl args = I"
    and args_n: "args_n args = n"
    and args_L: "args_L args = Formula.fv φ"
    and args_R: "args_R args = Formula.fv ψ"
    and args_pos: "args_pos args = pos"
  shows "wf_since_aux σ V R args φ ψ aux' (Suc ne)"
    and "qtable n (Formula.fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) ne (Sincep pos φ I ψ)) rel"
proof -
  let ?wf_tuple = "λv. wf_tuple n (Formula.fv ψ) v"
  note sat.simps[simp del]
  from pre[unfolded wf_since_aux_def] obtain cur auxlist where aux: "valid_msaux args cur aux auxlist"
    "sorted_wrt (λx y. fst y < fst x) auxlist"
    "t X. (t, X)  set auxlist  ne  0  t  τ σ (ne - 1)  τ σ (ne - 1) - t  right I 
      (i. τ σ i = t) 
      qtable n (fv ψ) (mem_restr R)
        (λv. Formula.sat σ V (map the v) (ne - 1) (Sincep pos φ (point (τ σ (ne - 1) - t)) ψ)) X"
    "t. ne  0  t  τ σ (ne - 1)  τ σ (ne - 1) - t  right I  (i. τ σ i = t) 
      (X. (t, X)  set auxlist)"
    and cur_def:
    "cur = (if ne = 0 then 0 else τ σ (ne - 1))"
    unfolding args_ivl args_n args_pos by blast
  from pre[unfolded wf_since_aux_def] have fv_sub: "Formula.fv φ  Formula.fv ψ" by simp

  define aux0 where "aux0 = join_msaux args rel1 (add_new_ts_msaux args (τ σ ne) aux)"
  define auxlist0 where "auxlist0 = [(t, join rel pos rel1). (t, rel)  auxlist, τ σ ne - t  right I]"
  have tabL: "table (args_n args) (args_L args) rel1"
    using qtable1[unfolded qtable_def] unfolding args_n[symmetric] args_L[symmetric] by simp
  have cur_le: "cur  τ σ ne"
    unfolding cur_def by auto
  have valid0: "valid_msaux args (τ σ ne) aux0 auxlist0" unfolding aux0_def auxlist0_def
    using valid_join_msaux[OF valid_add_new_ts_msaux[OF aux(1)], OF cur_le tabL]
    by (auto simp: args_ivl args_pos cur_def map_filter_alt split_beta cong: map_cong)
  from aux(2) have sorted_auxlist0: "sorted_wrt (λx y. fst x > fst y) auxlist0"
    unfolding auxlist0_def
    by (induction auxlist) (auto simp add: sorted_wrt_append)
  have in_auxlist0_1: "(t, X)  set auxlist0  ne  0  t  τ σ (ne-1)  τ σ ne - t  right I 
      (i. τ σ i = t) 
      qtable n (Formula.fv ψ) (mem_restr R) (λv. (Formula.sat σ V (map the v) (ne-1) (Sincep pos φ (point (τ σ (ne-1) - t)) ψ) 
        (if pos then Formula.sat σ V (map the v) ne φ else ¬ Formula.sat σ V (map the v) ne φ))) X" for t X
    unfolding auxlist0_def using fvi_subset
    by (auto 0 1 elim!: qtable_join[OF _ qtable1] simp: sat_the_restrict dest!: aux(3))
  then have in_auxlist0_le_τ: "(t, X)  set auxlist0  t  τ σ ne" for t X
    by (meson τ_mono diff_le_self le_trans)
  have in_auxlist0_2: "ne  0  t  τ σ (ne-1)  τ σ ne - t  right I  i. τ σ i = t 
    X. (t, X)  set auxlist0" for t
  proof -
    fix t
    assume "ne  0" "t  τ σ (ne-1)" "τ σ ne - t  right I" "i. τ σ i = t"
    then obtain X where "(t, X)  set auxlist"
      by (atomize_elim, intro aux(4))
        (auto simp: gr0_conv_Suc elim!: order_trans[rotated] intro!: diff_le_mono τ_mono)
    with τ σ ne - t  right I have "(t, join X pos rel1)  set auxlist0"
      unfolding auxlist0_def by (auto elim!: bexI[rotated] intro!: exI[of _ X])
    then show "X. (t, X)  set auxlist0"
      by blast
  qed
  have auxlist0_Nil: "auxlist0 = []  ne = 0  ne  0  (t. t  τ σ (ne-1)  τ σ ne - t  right I 
        (i. τ σ i = t))"
    using in_auxlist0_2 by (auto)

  have aux'_eq: "aux' = add_new_table_msaux args rel2 aux0"
    using result_eq unfolding aux0_def update_since_def Let_def by simp
  define auxlist' where
    auxlist'_eq: "auxlist' = (case auxlist0 of
      []  [(τ σ ne, rel2)]
    | x # auxlist'  (if fst x = τ σ ne then (fst x, snd x  rel2) # auxlist' else (τ σ ne, rel2) # x # auxlist'))"
  have tabR: "table (args_n args) (args_R args) rel2"
    using qtable2[unfolded qtable_def] unfolding args_n[symmetric] args_R[symmetric] by simp
  have valid': "valid_msaux args (τ σ ne) aux' auxlist'"
    unfolding aux'_eq auxlist'_eq using valid_add_new_table_msaux[OF valid0 tabR]
    by (auto simp: not_le split: list.splits option.splits if_splits)
  have sorted_auxlist': "sorted_wrt (λx y. fst x > fst y) auxlist'"
    unfolding auxlist'_eq
    using sorted_auxlist0 in_auxlist0_le_τ by (cases auxlist0) fastforce+
  have in_auxlist'_1: "t  τ σ ne  τ σ ne - t  right I  (i. τ σ i = t) 
      qtable n (Formula.fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) ne (Sincep pos φ (point (τ σ ne - t)) ψ)) X"
    if auxlist': "(t, X)  set auxlist'" for t X
  proof (cases auxlist0)
    case Nil
    with auxlist' show ?thesis
      unfolding auxlist'_eq using qtable2 auxlist0_Nil
      by (auto simp: zero_enat_def[symmetric] sat_Since_rec[where i=ne]
          dest: spec[of _ "τ σ (ne-1)"] elim!: qtable_cong[OF _ refl])
  next
    case (Cons a as)
    show ?thesis
    proof (cases "t = τ σ ne")
      case t: True
      show ?thesis
      proof (cases "fst a = τ σ ne")
        case True
        with auxlist' Cons t have "X = snd a  rel2"
          unfolding auxlist'_eq using sorted_auxlist0 by (auto split: if_splits)
        moreover from in_auxlist0_1[of "fst a" "snd a"] Cons have "ne  0"
          "fst a  τ σ (ne - 1)" "τ σ ne - fst a  right I" "i. τ σ i = fst a"
          "qtable n (fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) (ne - 1)
            (Sincep pos φ (point (τ σ (ne - 1) - fst a)) ψ)  (if pos then Formula.sat σ V (map the v) ne φ
              else ¬ Formula.sat σ V (map the v) ne φ)) (snd a)"
          by (auto simp: True[symmetric] zero_enat_def[symmetric])
        ultimately show ?thesis using qtable2 t True
          by (auto simp: sat_Since_rec[where i=ne] sat.simps(6) elim!: qtable_union)
      next
        case False
        with auxlist' Cons t have "X = rel2"
          unfolding auxlist'_eq using sorted_auxlist0 in_auxlist0_le_τ[of "fst a" "snd a"] by (auto split: if_splits)
        with auxlist' Cons t False show ?thesis
          unfolding auxlist'_eq using qtable2 in_auxlist0_2[of "τ σ (ne-1)"] in_auxlist0_le_τ[of "fst a" "snd a"] sorted_auxlist0
          by (auto simp: sat_Since_rec[where i=ne] sat.simps(3) zero_enat_def[symmetric] enat_0_iff not_le
              elim!: qtable_cong[OF _ refl] dest!: le_τ_less meta_mp)
      qed
    next
      case False
      with auxlist' Cons have "(t, X)  set auxlist0"
        unfolding auxlist'_eq by (auto split: if_splits)
      then have "ne  0" "t  τ σ (ne - 1)" "τ σ ne - t  right I" "i. τ σ i = t"
        "qtable n (fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) (ne - 1) (Sincep pos φ (point (τ σ (ne - 1) - t)) ψ) 
           (if pos then Formula.sat σ V (map the v) ne φ else ¬ Formula.sat σ V (map the v) ne φ)) X"
        using in_auxlist0_1 by blast+
      with False auxlist' Cons show ?thesis
        unfolding auxlist'_eq using qtable2
        by (fastforce simp: sat_Since_rec[where i=ne] sat.simps(6)
            diff_diff_right[where i="τ σ ne" and j="τ σ _ + τ σ ne" and k="τ σ (ne - 1)",
              OF trans_le_add2, simplified] elim!: qtable_cong[OF _ refl] order_trans dest: le_τ_less)
    qed
  qed

  have in_auxlist'_2: "X. (t, X)  set auxlist'" if "t  τ σ ne" "τ σ ne - t  right I" "i. τ σ i = t" for t
  proof (cases "t = τ σ ne")
    case True
    then show ?thesis
    proof (cases auxlist0)
      case Nil
      with True show ?thesis unfolding auxlist'_eq by (simp add: zero_enat_def[symmetric])
    next
      case (Cons a as)
      with True show ?thesis unfolding auxlist'_eq
        by (cases "fst a = τ σ ne") (auto simp: zero_enat_def[symmetric])
    qed
  next
    case False
    with that have "ne  0"
      using le_τ_less neq0_conv by blast
    moreover from False that have  "t  τ σ (ne-1)"
      by (metis One_nat_def Suc_leI Suc_pred τ_mono diff_is_0_eq' order.antisym neq0_conv not_le)
    ultimately obtain X where "(t, X)  set auxlist0" using τ σ ne - t  right I i. τ σ i = t
      using τ_mono[of "ne - 1" "ne" σ] by (atomize_elim, cases "right I") (auto intro!: in_auxlist0_2 simp del: τ_mono)
    then show ?thesis unfolding auxlist'_eq using False τ σ ne - t  right I
      by (auto intro: exI[of _ X] split: list.split)
  qed

  show "wf_since_aux σ V R args φ ψ aux' (Suc ne)"
    unfolding wf_since_aux_def args_ivl args_n args_pos
    by (auto simp add: fv_sub dest: in_auxlist'_1 intro: sorted_auxlist' in_auxlist'_2
        intro!: exI[of _ auxlist'] valid')

  have "rel = result_msaux args aux'"
    using result_eq by (auto simp add: update_since_def Let_def)
  with valid' have rel_eq: "rel = foldr (∪) [rel. (t, rel)  auxlist', left I  τ σ ne - t] {}"
    by (auto simp add: args_ivl valid_result_msaux
        intro!: arg_cong[where f = "λx. foldr (∪) (concat x) {}"] split: option.splits)
  have rel_alt: "rel = ((t, rel)  set auxlist'. if left I  τ σ ne - t then rel else empty_table)"
    unfolding rel_eq
    by (auto elim!: in_foldr_UnE bexI[rotated] intro!: in_foldr_UnI)
  show "qtable n (fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) ne (Sincep pos φ I ψ)) rel"
    unfolding rel_alt
  proof (rule qtable_Union[where Qi="λ(t, X) v.
    left I  τ σ ne - t  Formula.sat σ V (map the v) ne (Sincep pos φ (point (τ σ ne - t)) ψ)"],
        goal_cases finite qtable equiv)
    case (equiv v)
    show ?case
    proof (rule iffI, erule sat_Since_point, goal_cases left right)
      case (left j)
      then show ?case using in_auxlist'_2[of "τ σ j", OF _ _ exI, OF _ _ refl] by auto
    next
      case right
      then show ?case by (auto elim!: sat_Since_pointD dest: in_auxlist'_1)
    qed
  qed (auto dest!: in_auxlist'_1 intro!: qtable_empty)
qed

lemma fv_regex_from_mregex:
  "ok (length φs) mr  fv_regex (from_mregex mr φs)  (φ  set φs. fv φ)"
  by (induct mr) (auto simp: Bex_def in_set_conv_nth)+

lemma qtable_ε_lax:
  assumes "ok (length φs) mr"
    and "list_all2 (λφ rel. qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) i φ) rel) φs rels"
    and "fv_regex (from_mregex mr φs)  A" and "qtable n A (mem_restr R) Q guard"
  shows "qtable n A (mem_restr R)
   (λv. Regex.eps (Formula.sat σ V (map the v)) i (from_mregex mr φs)  Q v) (ε_lax guard rels mr)"
  using assms
proof (induct mr)
  case (MPlus mr1 mr2)
  from MPlus(3-6) show ?case
    by (auto intro!: qtable_union[OF MPlus(1,2)])
next
  case (MTimes mr1 mr2)
  then have "fv_regex (from_mregex mr1 φs)  A" "fv_regex (from_mregex mr2 φs)  A"
    using fv_regex_from_mregex[of φs mr1] fv_regex_from_mregex[of φs mr2] by (auto simp: subset_eq)
  with MTimes(3-6) show ?case
    by (auto simp: eps_the_restrict restrict_idle intro!: qtable_join[OF MTimes(1,2)])
qed (auto split: prod.splits if_splits simp: qtable_empty_iff list_all2_conv_all_nth
    in_set_conv_nth restrict_idle sat_the_restrict
    intro: in_qtableI qtableI elim!: qtable_join[where A=A and C=A])

lemma nullary_qtable_cases: "qtable n {} P Q X  (X = empty_table  X = unit_table n)"
  by (simp add: qtable_def table_empty)

lemma qtable_empty_unit_table:
  "qtable n {} R P empty_table  qtable n {} R (λv. ¬ P v) (unit_table n)"
  by (auto intro: qtable_unit_table simp add: qtable_empty_iff)

lemma qtable_unit_empty_table:
  "qtable n {} R P (unit_table n)  qtable n {} R (λv. ¬ P v) empty_table"
  by (auto intro!: qtable_empty elim: in_qtableE simp add: wf_tuple_empty unit_table_def)

lemma qtable_nonempty_empty_table:
  "qtable n {} R P X  x  X  qtable n {} R (λv. ¬ P v) empty_table"
  by (frule nullary_qtable_cases) (auto dest: qtable_unit_empty_table)


lemma qtable_rε_strict:
  assumes "safe_regex Past Strict (from_mregex mr φs)" "ok (length φs) mr" "A = fv_regex (from_mregex mr φs)"
    and "list_all2 (λφ rel. qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) i φ) rel) φs rels"
  shows "qtable n A (mem_restr R) (λv. Regex.eps (Formula.sat σ V (map the v)) i (from_mregex mr φs)) (rε_strict n rels mr)"
  using assms
proof (hypsubst, induct Past Strict "from_mregex mr φs" arbitrary: mr rule: safe_regex_induct)
  case (Skip n)
  then show ?case
    by (cases mr) (auto simp: qtable_empty_iff qtable_unit_table split: if_splits)
next
  case (Test φ)
  then show ?case
    by (cases mr) (auto simp: list_all2_conv_all_nth qtable_empty_unit_table
        dest!: qtable_nonempty_empty_table split: if_splits)
next
  case (Plus r s)
  then show ?case
    by (cases mr) (fastforce intro: qtable_union split: if_splits)+
next
  case (TimesP r s)
  then show ?case
    by (cases mr) (auto intro: qtable_cong[OF qtable_ε_lax] split: if_splits)+
next
  case (Star r)
  then show ?case
    by (cases mr) (auto simp: qtable_unit_table split: if_splits)
qed

lemma qtable_lε_strict:
  assumes "safe_regex Futu Strict (from_mregex mr φs)" "ok (length φs) mr" "A = fv_regex (from_mregex mr φs)"
    and "list_all2 (λφ rel. qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) i φ) rel) φs rels"
  shows "qtable n A (mem_restr R) (λv. Regex.eps (Formula.sat σ V (map the v)) i (from_mregex mr φs)) (lε_strict n rels mr)"
  using assms
proof (hypsubst, induct Futu Strict "from_mregex mr φs" arbitrary: mr rule: safe_regex_induct)
  case (Skip n)
  then show ?case
    by (cases mr) (auto simp: qtable_empty_iff qtable_unit_table split: if_splits)
next
  case (Test φ)
  then show ?case
    by (cases mr) (auto simp: list_all2_conv_all_nth qtable_empty_unit_table
        dest!: qtable_nonempty_empty_table split: if_splits)
next
  case (Plus r s)
  then show ?case
    by (cases mr) (fastforce intro: qtable_union split: if_splits)+
next
  case (TimesF r s)
  then show ?case
    by (cases mr) (auto intro: qtable_cong[OF qtable_ε_lax] split: if_splits)+
next
  case (Star r)
  then show ?case
    by (cases mr) (auto simp: qtable_unit_table split: if_splits)
qed

lemma rtranclp_False: "(λi j. False)** = (=)"
proof -
  have "(λi j. False)** i j  i = j" for i j :: 'a
    by (induct i j rule: rtranclp.induct) auto
  then show ?thesis
    by (auto intro: exI[of _ 0])
qed

inductive ok_rctxt for φs where
  "ok_rctxt φs id id"
| "ok_rctxt φs κ κ'  ok_rctxt φs (λt. κ (MTimes mr t)) (λt. κ' (Regex.Times (from_mregex mr φs) t))"

lemma ok_rctxt_swap: "ok_rctxt φs κ κ'  from_mregex (κ mr) φs = κ' (from_mregex mr φs)"
  by (induct κ κ' arbitrary: mr rule: ok_rctxt.induct) auto

lemma ok_rctxt_cong: "ok_rctxt φs κ κ'  Regex.match (Formula.sat σ V v) r = Regex.match (Formula.sat σ V v) s 
  Regex.match (Formula.sat σ V v) (κ' r) i j = Regex.match (Formula.sat σ V v) (κ' s) i j"
  by (induct κ κ' arbitrary: r s rule: ok_rctxt.induct) simp_all

lemma qtable_rδκ:
  assumes "ok (length φs) mr" "fv_regex (from_mregex mr φs)  A"
    and "list_all2 (λφ rel. qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) j φ) rel) φs rels"
    and "ok_rctxt φs κ κ'"
    and "ms  κ ` RPD mr. qtable n A (mem_restr R) (λv. Q (map the v) (from_mregex ms φs)) (lookup rel ms)"
  shows "qtable n A (mem_restr R)
  (λv. s  Regex.rpdκ κ' (Formula.sat σ V (map the v)) j (from_mregex mr φs). Q (map the v) s)
  ( κ rel rels mr)"
  using assms
proof (induct mr arbitrary: κ κ')
  case MSkip
  then show ?case
    by (auto simp: rtranclp_False ok_rctxt_swap qtable_empty_iff
        elim!: qtable_cong[OF _ _ ok_rctxt_cong[of _ κ κ']] split: nat.splits)
next
  case (MPlus mr1 mr2)
  from MPlus(3-7) show ?case
    by (auto intro!: qtable_union[OF MPlus(1,2)])
next
  case (MTimes mr1 mr2)
  from MTimes(3-7) show ?case
    by (auto intro!: qtable_union[OF MTimes(2) qtable_ε_lax[OF _ _ _ MTimes(1)]]
        elim!: ok_rctxt.intros(2) simp: MTimesL_def Ball_def)
next
  case (MStar mr)
  from MStar(2-6) show ?case
    by (auto intro!: qtable_cong[OF MStar(1)] intro: ok_rctxt.intros simp: MTimesL_def Ball_def)
qed (auto simp: qtable_empty_iff)

lemmas qtable_rδ = qtable_rδκ[OF _ _ _ ok_rctxt.intros(1), unfolded rpdκ_rpd image_id id_apply]

inductive ok_lctxt for φs where
  "ok_lctxt φs id id"
| "ok_lctxt φs κ κ'  ok_lctxt φs (λt. κ (MTimes t mr)) (λt. κ' (Regex.Times t (from_mregex mr φs)))"

lemma ok_lctxt_swap: "ok_lctxt φs κ κ'  from_mregex (κ mr) φs = κ' (from_mregex mr φs)"
  by (induct κ κ' arbitrary: mr rule: ok_lctxt.induct) auto

lemma ok_lctxt_cong: "ok_lctxt φs κ κ'  Regex.match (Formula.sat σ V v) r = Regex.match (Formula.sat σ V v) s 
  Regex.match (Formula.sat σ V v) (κ' r) i j = Regex.match (Formula.sat σ V v) (κ' s) i j"
  by (induct κ κ' arbitrary: r s rule: ok_lctxt.induct) simp_all

lemma qtable_lδκ:
  assumes "ok (length φs) mr" "fv_regex (from_mregex mr φs)  A"
    and "list_all2 (λφ rel. qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) j φ) rel) φs rels"
    and "ok_lctxt φs κ κ'"
    and "ms  κ ` LPD mr. qtable n A (mem_restr R) (λv. Q (map the v) (from_mregex ms φs)) (lookup rel ms)"
  shows "qtable n A (mem_restr R)
  (λv. s  Regex.lpdκ κ' (Formula.sat σ V (map the v)) j (from_mregex mr φs). Q (map the v) s)
  ( κ rel rels mr)"
  using assms
proof (induct mr arbitrary: κ κ')
  case MSkip
  then show ?case
    by (auto simp: rtranclp_False ok_lctxt_swap qtable_empty_iff
        elim!: qtable_cong[OF _ _ ok_rctxt_cong[of _ κ κ']] split: nat.splits)
next
  case (MPlus mr1 mr2)
  from MPlus(3-7) show ?case
    by (auto intro!: qtable_union[OF MPlus(1,2)])
next
  case (MTimes mr1 mr2)
  from MTimes(3-7) show ?case
    by (auto intro!: qtable_union[OF MTimes(1) qtable_ε_lax[OF _ _ _ MTimes(2)]]
        elim!: ok_lctxt.intros(2) simp: MTimesR_def Ball_def)
next
  case (MStar mr)
  from MStar(2-6) show ?case
    by (auto intro!: qtable_cong[OF MStar(1)] intro: ok_lctxt.intros simp: MTimesR_def Ball_def)
qed (auto simp: qtable_empty_iff)

lemmas qtable_lδ = qtable_lδκ[OF _ _ _ ok_lctxt.intros(1), unfolded lpdκ_lpd image_id id_apply]

lemma RPD_fv_regex_le:
  "ms  RPD mr  fv_regex (from_mregex ms φs)  fv_regex (from_mregex mr φs)"
  by (induct mr arbitrary: ms) (auto simp: MTimesL_def split: nat.splits)+

lemma RPD_safe: "safe_regex Past g (from_mregex mr φs) 
  ms  RPD mr  safe_regex Past g (from_mregex ms φs)"
proof (induct Past g "from_mregex mr φs" arbitrary: mr ms rule: safe_regex_induct)
  case Skip
  then show ?case
    by (cases mr) (auto split: nat.splits)
next
  case (Test g φ)
  then show ?case
    by (cases mr) auto
next
  case (Plus g r s mrs)
  then show ?case
  proof (cases mrs)
    case (MPlus mr ms)
    with Plus(3-5) show ?thesis
      by (auto dest!: Plus(1,2))
  qed auto
next
  case (TimesP g r s mrs)
  then show ?case
  proof (cases mrs)
    case (MTimes mr ms)
    with TimesP(3-5) show ?thesis
      by (cases g) (auto 0 4 simp: MTimesL_def dest: RPD_fv_regex_le TimesP(1,2))
  qed auto
next
  case (Star g r)
  then show ?case
  proof (cases mr)
    case (MStar x6)
    with Star(2-4) show ?thesis
      by (cases g) (auto 0 4 simp: MTimesL_def dest: RPD_fv_regex_le
          elim!: safe_cosafe[rotated] dest!: Star(1))
  qed auto
qed

lemma RPDi_safe: "safe_regex Past g (from_mregex mr φs) 
  ms  RPDi n mr ==> safe_regex Past g (from_mregex ms φs)"
  by (induct n arbitrary: ms mr) (auto dest: RPD_safe)

lemma RPDs_safe: "safe_regex Past g (from_mregex mr φs) 
  ms  RPDs mr ==> safe_regex Past g (from_mregex ms φs)"
  unfolding RPDs_def by (auto dest: RPDi_safe)

lemma RPD_safe_fv_regex: "safe_regex Past Strict (from_mregex mr φs) 
  ms  RPD mr  fv_regex (from_mregex ms φs) = fv_regex (from_mregex mr φs)"
proof (induct Past Strict "from_mregex mr φs" arbitrary: mr rule: safe_regex_induct)
  case (Skip n)
  then show ?case
    by (cases mr) (auto split: nat.splits)
next
  case (Test φ)
  then show ?case
    by (cases mr) auto
next
  case (Plus r s)
  then show ?case
    by (cases mr) auto
next
  case (TimesP r s)
  then show ?case
    by (cases mr) (auto 0 3 simp: MTimesL_def dest: RPD_fv_regex_le split: modality.splits)
next
  case (Star r)
  then show ?case
    by (cases mr) (auto 0 3 simp: MTimesL_def dest: RPD_fv_regex_le)
qed

lemma RPDi_safe_fv_regex: "safe_regex Past Strict (from_mregex mr φs) 
  ms  RPDi n mr ==> fv_regex (from_mregex ms φs) = fv_regex (from_mregex mr φs)"
  by (induct n arbitrary: ms mr) (auto 5 0 dest: RPD_safe_fv_regex RPD_safe)

lemma RPDs_safe_fv_regex: "safe_regex Past Strict (from_mregex mr φs) 
  ms  RPDs mr ==> fv_regex (from_mregex ms φs) = fv_regex (from_mregex mr φs)"
  unfolding RPDs_def by (auto dest: RPDi_safe_fv_regex)

lemma RPD_ok: "ok m mr  ms  RPD mr  ok m ms"
proof (induct mr arbitrary: ms)
  case (MPlus mr1 mr2)
  from MPlus(3,4) show ?case
    by (auto elim: MPlus(1,2))
next
  case (MTimes mr1 mr2)
  from MTimes(3,4) show ?case
    by (auto elim: MTimes(1,2) simp: MTimesL_def)
next
  case (MStar mr)
  from MStar(2,3) show ?case
    by (auto elim: MStar(1) simp: MTimesL_def)
qed (auto split: nat.splits)

lemma RPDi_ok: "ok m mr  ms  RPDi n mr  ok m ms"
  by (induct n arbitrary: ms mr) (auto intro: RPD_ok)

lemma RPDs_ok: "ok m mr  ms  RPDs mr  ok m ms"
  unfolding RPDs_def by (auto intro: RPDi_ok)

lemma LPD_fv_regex_le:
  "ms  LPD mr  fv_regex (from_mregex ms φs)  fv_regex (from_mregex mr φs)"
  by (induct mr arbitrary: ms) (auto simp: MTimesR_def split: nat.splits)+

lemma LPD_safe: "safe_regex Futu g (from_mregex mr φs) 
  ms  LPD mr ==> safe_regex Futu g (from_mregex ms φs)"
proof (induct Futu g "from_mregex mr φs" arbitrary: mr ms rule: safe_regex_induct)
  case Skip
  then show ?case
    by (cases mr) (auto split: nat.splits)
next
  case (Test g φ)
  then show ?case
    by (cases mr) auto
next
  case (Plus g r s mrs)
  then show ?case
  proof (cases mrs)
    case (MPlus mr ms)
    with Plus(3-5) show ?thesis
      by (auto dest!: Plus(1,2))
  qed auto
next
  case (TimesF g r s mrs)
  then show ?case
  proof (cases mrs)
    case (MTimes mr ms)
    with TimesF(3-5) show ?thesis
      by (cases g) (auto 0 4 simp: MTimesR_def dest: LPD_fv_regex_le split: modality.splits dest: TimesF(1,2))
  qed auto
next
  case (Star g r)
  then show ?case
  proof (cases mr)
    case (MStar x6)
    with Star(2-4) show ?thesis
      by (cases g) (auto 0 4 simp: MTimesR_def dest: LPD_fv_regex_le
          elim!: safe_cosafe[rotated] dest!: Star(1))
  qed auto
qed

lemma LPDi_safe: "safe_regex Futu g (from_mregex mr φs) 
  ms  LPDi n mr ==> safe_regex Futu g (from_mregex ms φs)"
  by (induct n arbitrary: ms mr) (auto dest: LPD_safe)

lemma LPDs_safe: "safe_regex Futu g (from_mregex mr φs) 
  ms  LPDs mr ==> safe_regex Futu g (from_mregex ms φs)"
  unfolding LPDs_def by (auto dest: LPDi_safe)

lemma LPD_safe_fv_regex: "safe_regex Futu Strict (from_mregex mr φs) 
  ms  LPD mr ==> fv_regex (from_mregex ms φs) = fv_regex (from_mregex mr φs)"
proof (induct Futu Strict "from_mregex mr φs" arbitrary: mr rule: safe_regex_induct)
  case Skip
  then show ?case
    by (cases mr) (auto split: nat.splits)
next
  case (Test φ)
  then show ?case
    by (cases mr) auto
next
  case (Plus r s)
  then show ?case
    by (cases mr) auto
next
  case (TimesF r s)
  then show ?case
    by (cases mr) (auto 0 3 simp: MTimesR_def dest: LPD_fv_regex_le split: modality.splits)
next
  case (Star r)
  then show ?case
    by (cases mr) (auto 0 3 simp: MTimesR_def dest: LPD_fv_regex_le)
qed

lemma LPDi_safe_fv_regex: "safe_regex Futu Strict (from_mregex mr φs) 
  ms  LPDi n mr ==> fv_regex (from_mregex ms φs) = fv_regex (from_mregex mr φs)"
  by (induct n arbitrary: ms mr) (auto 5 0 dest: LPD_safe_fv_regex LPD_safe)

lemma LPDs_safe_fv_regex: "safe_regex Futu Strict (from_mregex mr φs) 
  ms  LPDs mr ==> fv_regex (from_mregex ms φs) = fv_regex (from_mregex mr φs)"
  unfolding LPDs_def by (auto dest: LPDi_safe_fv_regex)

lemma LPD_ok: "ok m mr  ms  LPD mr  ok m ms"
proof (induct mr arbitrary: ms)
  case (MPlus mr1 mr2)
  from MPlus(3,4) show ?case
    by (auto elim: MPlus(1,2))
next
  case (MTimes mr1 mr2)
  from MTimes(3,4) show ?case
    by (auto elim: MTimes(1,2) simp: MTimesR_def)
next
  case (MStar mr)
  from MStar(2,3) show ?case
    by (auto elim: MStar(1) simp: MTimesR_def)
qed (auto split: nat.splits)

lemma LPDi_ok: "ok m mr  ms  LPDi n mr  ok m ms"
  by (induct n arbitrary: ms mr) (auto intro: LPD_ok)

lemma LPDs_ok: "ok m mr  ms  LPDs mr  ok m ms"
  unfolding LPDs_def by (auto intro: LPDi_ok)

lemma update_matchP:
  assumes pre: "wf_matchP_aux σ V n R I r aux ne"
    and safe: "safe_regex Past Strict r"
    and mr: "to_mregex r = (mr, φs)"
    and mrs: "mrs = sorted_list_of_set (RPDs mr)"
    and qtables: "list_all2 (λφ rel. qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) ne φ) rel) φs rels"
    and result_eq: "(rel, aux') = update_matchP n I mr mrs rels (τ σ ne) aux"
  shows "wf_matchP_aux σ V n R I r aux' (Suc ne)"
    and "qtable n (Formula.fv_regex r) (mem_restr R) (λv. Formula.sat σ V (map the v) ne (Formula.MatchP I r)) rel"
proof -
  let ?wf_tuple = "λv. wf_tuple n (Formula.fv_regex r) v"
  let ?update = "λrel t. mrtabulate mrs (λmr.
     id rel rels mr  (if t = τ σ ne then rε_strict n rels mr else {}))"
  note sat.simps[simp del]

  define aux0 where "aux0 = [(t, ?update rel t). (t, rel)  aux, enat (τ σ ne - t)  right I]"
  have sorted_aux0: "sorted_wrt (λx y. fst x > fst y) aux0"
    using pre[unfolded wf_matchP_aux_def, THEN conjunct1]
    unfolding aux0_def
    by (induction aux) (auto simp add: sorted_wrt_append)
  { fix ms
    assume "ms  RPDs mr"
    then have "fv_regex (from_mregex ms φs) = fv_regex r"
      "safe_regex Past Strict (from_mregex ms φs)" "ok (length φs) ms" "RPD ms  RPDs mr"
      using safe RPDs_safe RPDs_safe_fv_regex mr from_mregex_to_mregex RPDs_ok to_mregex_ok RPDs_trans
      by fastforce+
  } note * = this
  have **: "τ σ ne - (τ σ i + τ σ ne - τ σ (ne - Suc 0)) = τ σ (ne - Suc 0) - τ σ i" for i
    by (metis (no_types, lifting) Nat.diff_diff_right τ_mono add.commute add_diff_cancel_left diff_le_self le_add2 order_trans)
  have ***: "τ σ i = τ σ ne"
    if  "τ σ ne  τ σ i" "τ σ i  τ σ (ne - Suc 0)" "ne > 0" for i
    by (metis (no_types, lifting) Suc_pred τ_mono diff_le_self le_τ_less le_antisym not_less_eq that)
  then have in_aux0_1: "(t, X)  set aux0  ne  0  t  τ σ ne  τ σ ne - t  right I 
      (i. τ σ i = t) 
      (msRPDs mr. qtable n (fv_regex r) (mem_restr R) (λv. Formula.sat σ V (map the v) ne
         (Formula.MatchP (point (τ σ ne - t)) (from_mregex ms φs))) (lookup X ms))" for t X
    unfolding aux0_def using safe mr mrs
    by (auto simp: lookup_tabulate map_of_map_restrict restrict_map_def finite_RPDs * ** RPDs_trans diff_le_mono2
        intro!: sat_MatchP_rec[of σ _ _ ne, THEN iffD2]
        qtable_union[OF qtable_rδ[OF _ _ qtables] qtable_rε_strict[OF _ _ _ qtables],
          of ms "fv_regex r" "λv r. Formula.sat σ V v (ne - Suc 0) (Formula.MatchP (point 0) r)" _ ms for ms]
        qtable_cong[OF qtable_rδ[OF _ _ qtables],
          of ms "fv_regex r" "λv r. Formula.sat σ V v (ne - Suc 0) (Formula.MatchP (point (τ σ (ne - Suc 0) - τ σ i)) r)"
          _ _  "(λv. Formula.sat σ V (map the v) ne (Formula.MatchP (point (τ σ ne - τ σ i))  (from_mregex ms φs)))" for ms i]
        dest!: assms(1)[unfolded wf_matchP_aux_def, THEN conjunct2, THEN conjunct1, rule_format]
        sat_MatchP_rec["of" σ _ _ ne, THEN iffD1]
        elim!: bspec order.trans[OF _ τ_mono] bexI[rotated] split: option.splits if_splits) (* slow 7 sec *)
  then have in_aux0_le_τ: "(t, X)  set aux0  t  τ σ ne" for t X
    by (meson τ_mono diff_le_self le_trans)
  have in_aux0_2: "ne  0  t  τ σ (ne-1)  τ σ ne - t  right I  i. τ σ i = t 
    X. (t, X)  set aux0" for t
  proof -
    fix t
    assume "ne  0" "t  τ σ (ne-1)" "τ σ ne - t  right I" "i. τ σ i = t"
    then obtain X where "(t, X)  set aux"
      by (atomize_elim, intro assms(1)[unfolded wf_matchP_aux_def, THEN conjunct2, THEN conjunct2, rule_format])
        (auto simp: gr0_conv_Suc elim!: order_trans[rotated] intro!: diff_le_mono τ_mono)
    with τ σ ne - t  right I have "(t, ?update X t)  set aux0"
      unfolding aux0_def by (auto simp: id_def elim!: bexI[rotated] intro!: exI[of _ X])
    then show "X. (t, X)  set aux0"
      by blast
  qed
  have aux0_Nil: "aux0 = []  ne = 0  ne  0  (t. t  τ σ (ne-1)  τ σ ne - t  right I 
        (i. τ σ i = t))"
    using in_aux0_2 by (cases "ne = 0") (auto)

  have aux'_eq: "aux' = (case aux0 of
      []  [(τ σ ne, mrtabulate mrs (rε_strict n rels))]
    | x # aux'  (if fst x = τ σ ne then x # aux'
       else (τ σ ne, mrtabulate mrs (rε_strict n rels)) # x # aux'))"
    using result_eq unfolding aux0_def update_matchP_def Let_def by simp
  have sorted_aux': "sorted_wrt (λx y. fst x > fst y) aux'"
    unfolding aux'_eq
    using sorted_aux0 in_aux0_le_τ by (cases aux0) (fastforce)+

  have in_aux'_1: "t  τ σ ne  τ σ ne - t  right I  (i. τ σ i = t) 
     (msRPDs mr. qtable n (Formula.fv_regex r) (mem_restr R) (λv.
        Formula.sat σ V (map the v) ne (Formula.MatchP (point (τ σ ne - t)) (from_mregex ms φs))) (lookup X ms))"
    if aux': "(t, X)  set aux'" for t X
  proof (cases aux0)
    case Nil
    with aux' show ?thesis
      unfolding aux'_eq using safe mrs qtables aux0_Nil *
      by (auto simp: zero_enat_def[symmetric] sat_MatchP_rec[where i=ne]
          lookup_tabulate finite_RPDs split: option.splits
          intro!: qtable_cong[OF qtable_rε_strict]
          dest: spec[of _ "τ σ (ne-1)"])
  next
    case (Cons a as)
    show ?thesis
    proof (cases "t = τ σ ne")
      case t: True
      show ?thesis
      proof (cases "fst a = τ σ ne")
        case True
        with aux' Cons t have "X = snd a"
          unfolding aux'_eq using sorted_aux0 by auto
        moreover from in_aux0_1[of "fst a" "snd a"] Cons have "ne  0"
          "fst a  τ σ ne" "τ σ ne - fst a  right I" "i. τ σ i = fst a"
          "ms  RPDs mr. qtable n (fv_regex r) (mem_restr R) (λv. Formula.sat σ V (map the v) ne
            (Formula.MatchP (point (τ σ ne - fst a)) (from_mregex ms φs))) (lookup (snd a) ms)"
          by auto
        ultimately show ?thesis using t True
          by auto
      next
        case False
        with aux' Cons t have "X = mrtabulate mrs (rε_strict n rels)"
          unfolding aux'_eq using sorted_aux0 in_aux0_le_τ[of "fst a" "snd a"] by auto
        with aux' Cons t False show ?thesis
          unfolding aux'_eq using safe mrs qtables * in_aux0_2[of "τ σ (ne-1)"] in_aux0_le_τ[of "fst a" "snd a"] sorted_aux0
          by (auto simp: sat_MatchP_rec[where i=ne] zero_enat_def[symmetric] enat_0_iff not_le
              lookup_tabulate finite_RPDs split: option.splits
              intro!: qtable_cong[OF qtable_rε_strict] dest!: le_τ_less meta_mp)
      qed
    next
      case False
      with aux' Cons have "(t, X)  set aux0"
        unfolding aux'_eq by (auto split: if_splits)
      then have "ne  0" "t  τ σ ne" "τ σ ne - t  right I" "i. τ σ i = t"
        "ms  RPDs mr. qtable n (fv_regex r) (mem_restr R) (λv. Formula.sat σ V (map the v) ne
          (Formula.MatchP (point (τ σ ne - t)) (from_mregex ms φs))) (lookup X ms)"
        using in_aux0_1 by blast+
      with False aux' Cons show ?thesis
        unfolding aux'_eq by auto
    qed
  qed

  have in_aux'_2: "X. (t, X)  set aux'" if "t  τ σ ne" "τ σ ne - t  right I" "i. τ σ i = t" for t
  proof (cases "t = τ σ ne")
    case True
    then show ?thesis
    proof (cases aux0)
      case Nil
      with True show ?thesis unfolding aux'_eq by simp
    next
      case (Cons a as)
      with True show ?thesis unfolding aux'_eq using eq_fst_iff[of t a]
        by (cases "fst a = τ σ ne") auto
    qed
  next
    case False
    with that have "ne  0"
      using le_τ_less neq0_conv by blast
    moreover from False that have  "t  τ σ (ne-1)"
      by (metis One_nat_def Suc_leI Suc_pred τ_mono diff_is_0_eq' order.antisym neq0_conv not_le)
    ultimately obtain X where "(t, X)  set aux0" using τ σ ne - t  right I i. τ σ i = t
      by atomize_elim (auto intro!: in_aux0_2)
    then show ?thesis unfolding aux'_eq using False
      by (auto intro: exI[of _ X] split: list.split)
  qed

  show "wf_matchP_aux σ V n R I r aux' (Suc ne)"
    unfolding wf_matchP_aux_def using mr
    by (auto dest: in_aux'_1 intro: sorted_aux' in_aux'_2)

  have rel_eq: "rel = foldr (∪) [lookup rel mr. (t, rel)  aux', left I  τ σ ne - t] {}"
    unfolding aux'_eq aux0_def
    using result_eq by (simp add: update_matchP_def Let_def)
  have rel_alt: "rel = ((t, rel)  set aux'. if left I  τ σ ne - t then lookup rel mr else empty_table)"
    unfolding rel_eq
    by (auto elim!: in_foldr_UnE bexI[rotated] intro!: in_foldr_UnI)
  show "qtable n (fv_regex r) (mem_restr R) (λv. Formula.sat σ V (map the v) ne (Formula.MatchP I r)) rel"
    unfolding rel_alt
  proof (rule qtable_Union[where Qi="λ(t, X) v.
    left I  τ σ ne - t  Formula.sat σ V (map the v) ne (Formula.MatchP (point (τ σ ne - t)) r)"],
        goal_cases finite qtable equiv)
    case (equiv v)
    show ?case
    proof (rule iffI, erule sat_MatchP_point, goal_cases left right)
      case (left j)
      then show ?case using in_aux'_2[of "τ σ j", OF _ _ exI, OF _ _ refl] by auto
    next
      case right
      then show ?case by (auto elim!: sat_MatchP_pointD dest: in_aux'_1)
    qed
  qed (auto dest!: in_aux'_1 intro!: qtable_empty dest!: bspec[OF _ RPDs_refl]
      simp: from_mregex_eq[OF safe mr])
qed

lemma length_update_until: "length (update_until args rel1 rel2 nt aux) = Suc (length aux)"
  unfolding update_until_def by simp

lemma wf_update_until_auxlist:
  assumes pre: "wf_until_auxlist σ V n R pos φ I ψ auxlist ne"
    and qtable1: "qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) (ne + length auxlist) φ) rel1"
    and qtable2: "qtable n (Formula.fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) (ne + length auxlist) ψ) rel2"
    and fvi_subset: "Formula.fv φ  Formula.fv ψ"
    and args_ivl: "args_ivl args = I"
    and args_n: "args_n args = n"
    and args_pos: "args_pos args = pos"
  shows "wf_until_auxlist σ V n R pos φ I ψ (update_until args rel1 rel2 (τ σ (ne + length auxlist)) auxlist) ne"
  unfolding wf_until_auxlist_def length_update_until
  unfolding update_until_def list.rel_map add_Suc_right upt.simps eqTrueI[OF le_add1] if_True
proof (rule list_all2_appendI, unfold list.rel_map, goal_cases old new)
  case old
  show ?case
  proof (rule list.rel_mono_strong[OF assms(1)[unfolded wf_until_auxlist_def]]; safe, goal_cases mono1 mono2)
    case (mono1 i X Y v)
    then show ?case
      by (fastforce simp: args_ivl args_n args_pos sat_the_restrict less_Suc_eq
          elim!: qtable_join[OF _ qtable1] qtable_union[OF _ qtable1])
  next
    case (mono2 i X Y v)
    then show ?case using fvi_subset
      by (auto 0 3 simp: args_ivl args_n args_pos sat_the_restrict less_Suc_eq split: if_splits
          elim!: qtable_union[OF _ qtable_join_fixed[OF qtable2]]
          elim: qtable_cong[OF _ refl] intro: exI[of _ "ne + length auxlist"]) (* slow 8 sec*)
  qed
next
  case new
  then show ?case
    by (auto intro!: qtable_empty qtable1 qtable2[THEN qtable_cong] exI[of _ "ne + length auxlist"]
        simp: args_ivl args_n args_pos less_Suc_eq zero_enat_def[symmetric])
qed

lemma (in muaux) wf_update_until:
  assumes pre: "wf_until_aux σ V R args φ ψ aux ne"
    and qtable1: "qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) (ne + length_muaux args aux) φ) rel1"
    and qtable2: "qtable n (Formula.fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) (ne + length_muaux args aux) ψ) rel2"
    and fvi_subset: "Formula.fv φ  Formula.fv ψ"
    and args_ivl: "args_ivl args = I"
    and args_n: "args_n args = n"
    and args_L: "args_L args = Formula.fv φ"
    and args_R: "args_R args = Formula.fv ψ"
    and args_pos: "args_pos args = pos"
  shows "wf_until_aux σ V R args φ ψ (add_new_muaux args rel1 rel2 (τ σ (ne + length_muaux args aux)) aux) ne 
      length_muaux args (add_new_muaux args rel1 rel2 (τ σ (ne + length_muaux args aux)) aux) = Suc (length_muaux args aux)"
proof -
  from pre obtain cur auxlist where valid_aux: "valid_muaux args cur aux auxlist" and
    cur: "cur = (if ne + length auxlist = 0 then 0 else τ σ (ne + length auxlist - 1))" and
    pre_list: "wf_until_auxlist σ V n R pos φ I ψ auxlist ne"
    unfolding wf_until_aux_def args_ivl args_n args_pos by auto
  have length_aux: "length_muaux args aux = length auxlist"
    using valid_length_muaux[OF valid_aux] .
  define nt where "nt  τ σ (ne + length_muaux args aux)"
  have nt_mono: "cur  nt"
    unfolding cur nt_def length_aux by simp
  define auxlist' where "auxlist'  update_until args rel1 rel2 (τ σ (ne + length auxlist)) auxlist"
  have length_auxlist': "length auxlist' = Suc (length auxlist)"
    unfolding auxlist'_def by (auto simp add: length_update_until)
  have tab1: "table (args_n args) (args_L args) rel1"
    using qtable1 unfolding args_n[symmetric] args_L[symmetric] by (auto simp add: qtable_def)
  have tab2: "table (args_n args) (args_R args) rel2"
    using qtable2 unfolding args_n[symmetric] args_R[symmetric] by (auto simp add: qtable_def)
  have fv_sub: "fv φ  fv ψ"
    using pre unfolding wf_until_aux_def by auto
  moreover have valid_add_new_auxlist: "valid_muaux args nt (add_new_muaux args rel1 rel2 nt aux) auxlist'"
    using valid_add_new_muaux[OF valid_aux tab1 tab2 nt_mono]
    unfolding auxlist'_def nt_def length_aux .
  moreover have "length_muaux args (add_new_muaux args rel1 rel2 nt aux) = Suc (length_muaux args aux)"
    using valid_length_muaux[OF valid_add_new_auxlist] unfolding length_auxlist' length_aux[symmetric] .
  moreover have "wf_until_auxlist σ V n R pos φ I ψ auxlist' ne"
    using wf_update_until_auxlist[OF pre_list qtable1[unfolded length_aux] qtable2[unfolded length_aux] fv_sub args_ivl args_n args_pos]
    unfolding auxlist'_def .
  moreover have "τ σ (ne + length auxlist) = (if ne + length auxlist' = 0 then 0 else τ σ (ne + length auxlist' - 1))"
    unfolding cur length_auxlist' by auto
  ultimately show ?thesis
    unfolding wf_until_aux_def nt_def length_aux args_ivl args_n args_pos by fast
qed

lemma length_update_matchF_base:
  "length (fst (update_matchF_base I mr mrs nt entry st)) = Suc 0"
  by (auto simp: Let_def update_matchF_base_def)

lemma length_update_matchF_step:
  "length (fst (update_matchF_step I mr mrs nt entry st)) = Suc (length (fst st))"
  by (auto simp: Let_def update_matchF_step_def split: prod.splits)

lemma length_foldr_update_matchF_step:
  "length (fst (foldr (update_matchF_step I mr mrs nt) aux base)) = length aux + length (fst base)"
  by (induct aux arbitrary: base) (auto simp: Let_def length_update_matchF_step)

lemma length_update_matchF: "length (update_matchF n I mr mrs rels nt aux) = Suc (length aux)"
  unfolding update_matchF_def update_matchF_base_def length_foldr_update_matchF_step
  by (auto simp: Let_def)

lemma wf_update_matchF_base_invar:
  assumes safe: "safe_regex Futu Strict r"
    and mr: "to_mregex r = (mr, φs)"
    and mrs: "mrs = sorted_list_of_set (LPDs mr)"
    and qtables: "list_all2 (λφ rel. qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) j φ) rel) φs rels"
  shows "wf_matchF_invar σ V n R I r (update_matchF_base n I mr mrs rels (τ σ j)) j"
proof -
  have from_mregex: "from_mregex mr φs = r"
    using safe mr using from_mregex_eq by blast
  { fix ms
    assume "ms  LPDs mr"
    then have "fv_regex (from_mregex ms φs) = fv_regex r"
      "safe_regex Futu Strict (from_mregex ms φs)" "ok (length φs) ms" "LPD ms  LPDs mr"
      using safe LPDs_safe LPDs_safe_fv_regex mr from_mregex_to_mregex LPDs_ok to_mregex_ok LPDs_trans
      by fastforce+
  } note * = this
  show ?thesis
    unfolding wf_matchF_invar_def wf_matchF_aux_def update_matchF_base_def mr prod.case Let_def mrs
    using safe
    by (auto simp: * from_mregex qtables qtable_empty_iff zero_enat_def[symmetric]
        lookup_tabulate finite_LPDs eps_match less_Suc_eq LPDs_refl
        intro!: qtable_cong[OF qtable_lε_strict[where φs=φs]] intro: qtables exI[of _ j]
        split: option.splits)
qed

lemma Un_empty_table[simp]: "rel  empty_table = rel" "empty_table  rel = rel"
  unfolding empty_table_def by auto

lemma wf_matchF_invar_step:
  assumes wf: "wf_matchF_invar σ V n R I r st (Suc i)"
    and safe: "safe_regex Futu Strict r"
    and mr: "to_mregex r = (mr, φs)"
    and mrs: "mrs = sorted_list_of_set (LPDs mr)"
    and qtables: "list_all2 (λφ rel. qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) i φ) rel) φs rels"
    and rel: "qtable n (Formula.fv_regex r) (mem_restr R) (λv. (j. i  j  j < i + length (fst st)  mem (τ σ j - τ σ i) I 
          Regex.match (Formula.sat σ V (map the v)) r i j)) rel"
    and entry: "entry = (τ σ i, rels, rel)"
    and nt: "nt = τ σ (i + length (fst st))"
  shows "wf_matchF_invar σ V n R I r (update_matchF_step I mr mrs nt entry st) i"
proof -
  have from_mregex: "from_mregex mr φs = r"
    using safe mr using from_mregex_eq by blast
  { fix ms
    assume "ms  LPDs mr"
    then have "fv_regex (from_mregex ms φs) = fv_regex r"
      "safe_regex Futu Strict (from_mregex ms φs)" "ok (length φs) ms" "LPD ms  LPDs mr"
      using safe LPDs_safe LPDs_safe_fv_regex mr from_mregex_to_mregex LPDs_ok to_mregex_ok LPDs_trans
      by fastforce+
  } note * = this
  { fix aux X ms
    assume "st = (aux, X)" "ms  LPDs mr"
    with wf mr have "qtable n (fv_regex r) (mem_restr R)
      (λv. Regex.match (Formula.sat σ V (map the v)) (from_mregex ms φs) i (i + length aux))
      ( (λx. x) X rels ms)"
      by (intro qtable_cong[OF qtable_lδ[where φs=φs and A="fv_regex r" and
              Q="λv r. Regex.match (Formula.sat σ V v) r (Suc i) (i + length aux)", OF _ _ qtables]])
        (auto simp: wf_matchF_invar_def * LPDs_trans lpd_match[of i] elim!: bspec)
  } note  = this
  have "lookup (mrtabulate mrs f) ms = f ms" if "ms  LPDs mr" for ms and f :: "mregex  'a table"
    using that mrs  by (fastforce simp: lookup_tabulate finite_LPDs split: option.splits)+
  then show ?thesis
    using wf mr mrs entry nt LPDs_trans
    by (auto 0 3 simp: Let_def wf_matchF_invar_def update_matchF_step_def wf_matchF_aux_def mr * LPDs_refl
        list_all2_Cons1 append_eq_Cons_conv upt_eq_Cons_conv Suc_le_eq qtables
        lookup_tabulate finite_LPDs id_def  from_mregex less_Suc_eq
        intro!: qtable_union[OF rel ] qtable_cong[OF rel]
        intro: exI[of _ "i + length _"]
        split: if_splits prod.splits)
qed

lemma wf_update_matchF_invar:
  assumes pre: "wf_matchF_aux σ V n R I r aux ne (length (fst st) - 1)"
    and wf: "wf_matchF_invar σ V n R I r st (ne + length aux)"
    and safe: "safe_regex Futu Strict r"
    and mr: "to_mregex r = (mr, φs)"
    and mrs: "mrs = sorted_list_of_set (LPDs mr)"
    and j: "j = ne + length aux + length (fst st) - 1"
  shows "wf_matchF_invar σ V n R I r (foldr (update_matchF_step I mr mrs (τ σ j)) aux st) ne"
  using pre wf unfolding j
proof (induct aux arbitrary: ne)
  case (Cons entry aux)
  from Cons(1)[of "Suc ne"] Cons(2,3) show ?case
    unfolding foldr.simps o_apply
    by (intro wf_matchF_invar_step[where rels = "fst (snd entry)" and rel = "snd (snd entry)"])
      (auto simp: safe mr mrs wf_matchF_aux_def wf_matchF_invar_def list_all2_Cons1 append_eq_Cons_conv
        Suc_le_eq upt_eq_Cons_conv length_foldr_update_matchF_step add.assoc split: if_splits)
qed simp


lemma wf_update_matchF:
  assumes pre: "wf_matchF_aux σ V n R I r aux ne 0"
    and safe: "safe_regex Futu Strict r"
    and mr: "to_mregex r = (mr, φs)"
    and mrs: "mrs = sorted_list_of_set (LPDs mr)"
    and nt: "nt = τ σ (ne + length aux)"
    and qtables: "list_all2 (λφ rel. qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) (ne + length aux) φ) rel) φs rels"
  shows "wf_matchF_aux σ V n R I r (update_matchF n I mr mrs rels nt aux) ne 0"
  unfolding update_matchF_def using wf_update_matchF_base_invar[OF safe mr mrs qtables, of I]
  unfolding nt
  by (intro wf_update_matchF_invar[OF _ _ safe mr mrs, unfolded wf_matchF_invar_def split_beta, THEN conjunct2, THEN conjunct1])
    (auto simp: length_update_matchF_base wf_matchF_invar_def update_matchF_base_def Let_def pre)

lemma wf_until_aux_Cons: "wf_until_auxlist σ V n R pos φ I ψ (a # aux) ne 
  wf_until_auxlist σ V n R pos φ I ψ aux (Suc ne)"
  unfolding wf_until_auxlist_def
  by (simp add: upt_conv_Cons del: upt_Suc cong: if_cong)

lemma wf_matchF_aux_Cons: "wf_matchF_aux σ V n R I r (entry # aux) ne i 
  wf_matchF_aux σ V n R I r aux (Suc ne) i"
  unfolding wf_matchF_aux_def
  by (simp add: upt_conv_Cons del: upt_Suc cong: if_cong split: prod.splits)

lemma wf_until_aux_Cons1: "wf_until_auxlist σ V n R pos φ I ψ ((t, a1, a2) # aux) ne  t = τ σ ne"
  unfolding wf_until_auxlist_def
  by (simp add: upt_conv_Cons del: upt_Suc)

lemma wf_matchF_aux_Cons1: "wf_matchF_aux σ V n R I r ((t, rels, rel) # aux) ne i  t = τ σ ne"
  unfolding wf_matchF_aux_def
  by (simp add: upt_conv_Cons del: upt_Suc split: prod.splits)

lemma wf_until_aux_Cons3: "wf_until_auxlist σ V n R pos φ I ψ ((t, a1, a2) # aux) ne 
  qtable n (Formula.fv ψ) (mem_restr R) (λv. (j. ne  j  j < Suc (ne + length aux)  mem (τ σ j - τ σ ne) I 
    Formula.sat σ V (map the v) j ψ  (k{ne..<j}. if pos then Formula.sat σ V (map the v) k φ else ¬ Formula.sat σ V (map the v) k φ))) a2"
  unfolding wf_until_auxlist_def
  by (simp add: upt_conv_Cons del: upt_Suc)

lemma wf_matchF_aux_Cons3: "wf_matchF_aux σ V n R I r ((t, rels, rel) # aux) ne i 
  qtable n (Formula.fv_regex r) (mem_restr R) (λv. (j. ne  j  j < Suc (ne + length aux + i)  mem (τ σ j - τ σ ne) I 
    Regex.match (Formula.sat σ V (map the v)) r ne j)) rel"
  unfolding wf_matchF_aux_def
  by (simp add: upt_conv_Cons del: upt_Suc split: prod.splits)

lemma upt_append: "a  b  b  c  [a..<b] @ [b..<c] = [a..<c]"
  by (metis le_Suc_ex upt_add_eq_append)

lemma wf_mbuf2_add:
  assumes "wf_mbuf2 i ja jb P Q buf"
    and "list_all2 P [ja..<ja'] xs"
    and "list_all2 Q [jb..<jb'] ys"
    and "ja  ja'" "jb  jb'"
  shows "wf_mbuf2 i ja' jb' P Q (mbuf2_add xs ys buf)"
  using assms unfolding wf_mbuf2_def
  by (auto 0 3 simp: list_all2_append2 upt_append dest: list_all2_lengthD
      intro: exI[where x="[i..<ja]"] exI[where x="[ja..<ja']"]
      exI[where x="[i..<jb]"] exI[where x="[jb..<jb']"] split: prod.splits)

lemma wf_mbufn_add:
  assumes "wf_mbufn i js Ps buf"
    and "list_all3 list_all2 Ps (List.map2 (λj j'. [j..<j']) js js') xss"
    and "list_all2 (≤) js js'"
  shows "wf_mbufn i js' Ps (mbufn_add xss buf)"
  unfolding wf_mbufn_def list_all3_conv_all_nth
proof safe
  show "length Ps = length js'" "length js' = length (mbufn_add xss buf)"
    using assms unfolding wf_mbufn_def list_all3_conv_all_nth list_all2_conv_all_nth by auto
next
  fix k assume k: "k < length Ps"
  then show "i  js' ! k"
    using assms unfolding wf_mbufn_def list_all3_conv_all_nth list_all2_conv_all_nth
    by (auto 0 4 dest: spec[of _ i])
  from k have " [i..<js' ! k] =  [i..<js ! k] @ [js ! k ..<js' ! k]" and
    "length [i..<js ! k] = length (buf ! k)"
    using assms(1,3) unfolding wf_mbufn_def list_all3_conv_all_nth list_all2_conv_all_nth
    by (auto simp: upt_append)
  with k show "list_all2 (Ps ! k) [i..<js' ! k] (mbufn_add xss buf ! k)"
    using assms[unfolded wf_mbufn_def list_all3_conv_all_nth]
    by (auto simp add: list_all2_append)
qed

lemma mbuf2_take_eqD:
  assumes "mbuf2_take f buf = (xs, buf')"
    and "wf_mbuf2 i ja jb P Q buf"
  shows "wf_mbuf2 (min ja jb) ja jb P Q buf'"
    and "list_all2 (λi z. x y. P i x  Q i y  z = f x y) [i..<min ja jb] xs"
  using assms unfolding wf_mbuf2_def
  by (induction f buf arbitrary: i xs buf' rule: mbuf2_take.induct)
    (fastforce simp add: list_all2_Cons2 upt_eq_Cons_conv min_absorb1 min_absorb2 split: prod.splits)+

lemma list_all3_Nil[simp]:
  "list_all3 P xs ys []  xs = []  ys = []"
  "list_all3 P xs [] zs  xs = []  zs = []"
  "list_all3 P [] ys zs  ys = []  zs = []"
  unfolding list_all3_conv_all_nth by auto

lemma list_all3_Cons:
  "list_all3 P xs ys (z # zs)  (x xs' y ys'. xs = x # xs'  ys = y # ys'  P x y z  list_all3 P xs' ys' zs)"
  "list_all3 P xs (y # ys) zs  (x xs' z zs'. xs = x # xs'  zs = z # zs'  P x y z  list_all3 P xs' ys zs')"
  "list_all3 P (x # xs) ys zs  (y ys' z zs'. ys = y # ys'  zs = z # zs'  P x y z  list_all3 P xs ys' zs')"
  unfolding list_all3_conv_all_nth
  by (auto simp: length_Suc_conv Suc_length_conv nth_Cons split: nat.splits)

lemma list_all3_mono_strong: "list_all3 P xs ys zs 
  (x y z. x  set xs  y  set ys  z  set zs  P x y z  Q x y z) 
  list_all3 Q xs ys zs"
  by (induct xs ys zs rule: list_all3.induct) (auto intro: list_all3.intros)

definition Mini where
  "Mini i js = (if js = [] then i else Min (set js))"

lemma wf_mbufn_in_set_Mini:
  assumes "wf_mbufn i js Ps buf"
  shows "[]  set buf  Mini i js = i"
  unfolding in_set_conv_nth
proof (elim exE conjE)
  fix k
  have "i  j" if "j  set js" for j
    using that assms unfolding wf_mbufn_def list_all3_conv_all_nth in_set_conv_nth by auto
  moreover assume "k < length buf" "buf ! k = []"
  ultimately show ?thesis using assms
    unfolding Mini_def wf_mbufn_def list_all3_conv_all_nth
    by (auto 0 4 dest!: spec[of _ k] intro: Min_eqI simp: in_set_conv_nth)
qed

lemma wf_mbufn_notin_set:
  assumes "wf_mbufn i js Ps buf"
  shows "[]  set buf  j  set js  i < j"
  using assms unfolding wf_mbufn_def list_all3_conv_all_nth
  by (cases "i  set js") (auto intro: le_neq_implies_less simp: in_set_conv_nth)

lemma wf_mbufn_map_tl:
  "wf_mbufn i js Ps buf  []  set buf  wf_mbufn (Suc i) js Ps (map tl buf)"
  by (auto simp: wf_mbufn_def list_all3_map Suc_le_eq
      dest: rel_funD[OF tl_transfer]  elim!: list_all3_mono_strong le_neq_implies_less)

lemma list_all3_list_all2I: "list_all3 (λx y z. Q x z) xs ys zs  list_all2 Q xs zs"
  by (induct xs ys zs rule: list_all3.induct) auto

lemma mbuf2t_take_eqD:
  assumes "mbuf2t_take f z buf nts = (z', buf', nts')"
    and "wf_mbuf2 i ja jb P Q buf"
    and "list_all2 R [i..<j] nts"
    and "ja  j" "jb  j"
  shows "wf_mbuf2 (min ja jb) ja jb P Q buf'"
    and "list_all2 R [min ja jb..<j] nts'"
  using assms unfolding wf_mbuf2_def
  by (induction f z buf nts arbitrary: i z' buf' nts' rule: mbuf2t_take.induct)
    (auto simp add: list_all2_Cons2 upt_eq_Cons_conv less_eq_Suc_le min_absorb1 min_absorb2
      split: prod.split)

lemma wf_mbufn_take:
  assumes "mbufn_take f z buf = (z', buf')"
    and "wf_mbufn i js Ps buf"
  shows "wf_mbufn (Mini i js) js Ps buf'"
  using assms unfolding wf_mbufn_def
proof (induction f z buf arbitrary: i z' buf' rule: mbufn_take.induct)
  case rec: (1 f z buf)
  show ?case proof (cases "buf = []")
    case True
    with rec.prems show ?thesis by simp
  next
    case nonempty: False
    show ?thesis proof (cases "[]  set buf")
      case True
      from rec.prems(2) have "jset js. i  j"
        by (auto simp: in_set_conv_nth list_all3_conv_all_nth)
      moreover from True rec.prems(2) have "i  set js"
        by (fastforce simp: in_set_conv_nth list_all3_conv_all_nth list_all2_iff)
      ultimately have "Mini i js = i"
        unfolding Mini_def
        by (auto intro!: antisym[OF Min.coboundedI Min.boundedI])
      with rec.prems nonempty True show ?thesis by simp
    next
      case False
      from nonempty rec.prems(2) have "Mini i js = Mini (Suc i) js"
        unfolding Mini_def by auto
      show ?thesis
        unfolding Mini i js = Mini (Suc i) js
      proof (rule rec.IH)
        show "¬ (buf = []  []  set buf)" using nonempty False by simp
        show "list_all3 (λP j xs. Suc i  j  list_all2 P [Suc i..<j] xs) Ps js (map tl buf)"
          using False rec.prems(2)
          by (auto simp: list_all3_map elim!: list_all3_mono_strong dest: list.rel_sel[THEN iffD1])
        show "mbufn_take f (f (map hd buf) z) (map tl buf) = (z', buf')"
          using nonempty False rec.prems(1) by simp
      qed
    qed
  qed
qed

lemma mbufnt_take_eqD:
  assumes "mbufnt_take f z buf nts = (z', buf', nts')"
    and "wf_mbufn i js Ps buf"
    and "list_all2 R [i..<j] nts"
    and "k. k  set js  k  j"
    and "k = Mini (i + length nts) js"
  shows "wf_mbufn k js Ps buf'"
    and "list_all2 R [k..<j] nts'"
  using assms(1-4) unfolding assms(5)
proof (induction f z buf nts arbitrary: i z' buf' nts' rule: mbufnt_take.induct)
  case IH: (1 f z buf nts)
  note mbufnt_take.simps[simp del]
  case 1
  then have *: "list_all2 R [Suc i..<j] (tl nts)"
    by (auto simp: list.rel_sel[of R "[i..<j]" nts, simplified])
  from 1 show ?case
    using wf_mbufn_in_set_Mini[OF 1(2)]
    by (subst (asm) mbufnt_take.simps)
      (force simp: Mini_def wf_mbufn_def split: if_splits prod.splits elim!: list_all3_mono_strong
        dest!: IH(1)[rotated, OF _ wf_mbufn_map_tl[OF 1(2)] *])
  case 2
  then have *: "list_all2 R [Suc i..<j] (tl nts)"
    by (auto simp: list.rel_sel[of R "[i..<j]" nts, simplified])
  have [simp]: "Suc (i + (length nts - Suc 0)) = i + length nts" if "nts  []"
    using that by (fastforce simp flip: length_greater_0_conv)
  with 2 show ?case
    using wf_mbufn_in_set_Mini[OF 2(2)] wf_mbufn_notin_set[OF 2(2)]
    by (subst (asm) mbufnt_take.simps) (force simp: Mini_def wf_mbufn_def
        dest!: IH(2)[rotated, OF _ wf_mbufn_map_tl[OF 2(2)] *]
        split: if_splits prod.splits)
qed

lemma mbuf2t_take_induct[consumes 5, case_names base step]:
  assumes "mbuf2t_take f z buf nts = (z', buf', nts')"
    and "wf_mbuf2 i ja jb P Q buf"
    and "list_all2 R [i..<j] nts"
    and "ja  j" "jb  j"
    and "U i z"
    and "k x y t z. i  k  Suc k  ja  Suc k  jb 
      P k x  Q k y  R k t  U k z  U (Suc k) (f x y t z)"
  shows "U (min ja jb) z'"
  using assms unfolding wf_mbuf2_def
  by (induction f z buf nts arbitrary: i z' buf' nts' rule: mbuf2t_take.induct)
    (auto simp add: list_all2_Cons2 upt_eq_Cons_conv less_eq_Suc_le min_absorb1 min_absorb2
      elim!: arg_cong2[of _ _ _ _ U, OF _ refl, THEN iffD1, rotated] split: prod.split)

lemma list_all2_hdD:
  assumes "list_all2 P [i..<j] xs" "xs  []"
  shows "P i (hd xs)" "i < j"
  using assms unfolding list_all2_conv_all_nth
  by (auto simp: hd_conv_nth intro: zero_less_diff[THEN iffD1] dest!: spec[of _ 0])

lemma mbufn_take_induct[consumes 3, case_names base step]:
  assumes "mbufn_take f z buf = (z', buf')"
    and "wf_mbufn i js Ps buf"
    and "U i z"
    and "k xs z. i  k  Suc k  Mini i js 
      list_all2 (λP x. P k x) Ps xs  U k z  U (Suc k) (f xs z)"
  shows "U (Mini i js) z'"
  using assms unfolding wf_mbufn_def
proof (induction f z buf arbitrary: i z' buf' rule: mbufn_take.induct)
  case rec: (1 f z buf)
  show ?case proof (cases "buf = []")
    case True
    with rec.prems show ?thesis unfolding Mini_def by simp
  next
    case nonempty: False
    show ?thesis proof (cases "[]  set buf")
      case True
      from rec.prems(2) have "jset js. i  j"
        by (auto simp: in_set_conv_nth list_all3_conv_all_nth)
      moreover from True rec.prems(2) have "i  set js"
        by (fastforce simp: in_set_conv_nth list_all3_conv_all_nth list_all2_iff)
      ultimately have "Mini i js = i"
        unfolding Mini_def
        by (auto intro!: antisym[OF Min.coboundedI Min.boundedI])
      with rec.prems nonempty True show ?thesis by simp
    next
      case False
      with nonempty rec.prems(2) have more: "Suc i  Mini i js"
        using diff_is_0_eq not_le unfolding Mini_def
        by (fastforce simp: in_set_conv_nth list_all3_conv_all_nth list_all2_iff)
      then have "Mini i js = Mini (Suc i) js" unfolding Mini_def by auto
      show ?thesis
        unfolding Mini i js = Mini (Suc i) js
      proof (rule rec.IH)
        show "¬ (buf = []  []  set buf)" using nonempty False by simp
        show "mbufn_take f (f (map hd buf) z) (map tl buf) = (z', buf')"
          using nonempty False rec.prems by simp
        show "list_all3 (λP j xs. Suc i  j  list_all2 P [Suc i..<j] xs) Ps js (map tl buf)"
          using False rec.prems
          by (auto simp: list_all3_map elim!: list_all3_mono_strong dest: list.rel_sel[THEN iffD1])
        show "U (Suc i) (f (map hd buf) z)"
          using more False rec.prems
          by (auto 0 4 simp: list_all3_map intro!: rec.prems(4) list_all3_list_all2I
              elim!: list_all3_mono_strong dest: list.rel_sel[THEN iffD1])
        show "k xs z. Suc i  k  Suc k  Mini (Suc i) js 
          list_all2 (λP. P k) Ps xs  U k z  U (Suc k) (f xs z)"
          by (rule rec.prems(4)) (auto simp: Mini_def)
      qed
    qed
  qed
qed

lemma mbufnt_take_induct[consumes 5, case_names base step]:
  assumes "mbufnt_take f z buf nts = (z', buf', nts')"
    and "wf_mbufn i js Ps buf"
    and "list_all2 R [i..<j] nts"
    and "k. k  set js  k  j"
    and "U i z"
    and "k xs t z. i  k  Suc k  Mini j js 
      list_all2 (λP x. P k x) Ps xs  R k t  U k z  U (Suc k) (f xs t z)"
  shows "U (Mini (i + length nts) js) z'"
  using assms
proof (induction f z buf nts arbitrary: i z' buf' nts' rule: mbufnt_take.induct)
  case (1 f z buf nts)
  then have *: "list_all2 R [Suc i..<j] (tl nts)"
    by (auto simp: list.rel_sel[of R "[i..<j]" nts, simplified])
  note mbufnt_take.simps[simp del]
  from 1(2-6) have "i = Min (set js)" if "js  []" "nts = []"
    using that unfolding wf_mbufn_def using wf_mbufn_in_set_Mini[OF 1(3)]
    by (fastforce simp: Mini_def list_all3_Cons neq_Nil_conv)
  with 1(2-7) list_all2_hdD[OF 1(4)] show ?case
    unfolding wf_mbufn_def using wf_mbufn_in_set_Mini[OF 1(3)] wf_mbufn_notin_set[OF 1(3)]
    by (subst (asm) mbufnt_take.simps)
      (auto simp add: Mini_def list.rel_map Suc_le_eq
        elim!: arg_cong2[of _ _ _ _ U, OF _ refl, THEN iffD1, rotated]
        list_all3_mono_strong[THEN list_all3_list_all2I[of _ _ js]] list_all2_hdD
        dest!: 1(1)[rotated, OF _ wf_mbufn_map_tl[OF 1(3)] * _ 1(7)] split: prod.split if_splits)
qed

lemma mbuf2_take_add':
  assumes eq: "mbuf2_take f (mbuf2_add xs ys buf) = (zs, buf')"
    and pre: "wf_mbuf2' σ P V j n R φ ψ buf"
    and rm: "rel_mapping (≤) P P'"
    and xs: "list_all2 (λi. qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) i φ))
      [progress σ P φ j..<progress σ P' φ j'] xs"
    and ys: "list_all2 (λi. qtable n (Formula.fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) i ψ))
      [progress σ P ψ j..<progress σ P' ψ j'] ys"
    and "j  j'"
  shows "wf_mbuf2' σ P' V j' n R φ ψ buf'"
    and "list_all2 (λi Z. X Y.
      qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) i φ) X 
      qtable n (Formula.fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) i ψ) Y 
      Z = f X Y)
      [min (progress σ P φ j) (progress σ P ψ j)..<min (progress σ P' φ j') (progress σ P' ψ j')] zs"
  using pre rm unfolding wf_mbuf2'_def
  by (force intro!: mbuf2_take_eqD[OF eq] wf_mbuf2_add[OF _ xs ys] progress_mono_gen[OF j  j' rm])+

lemma mbuf2t_take_add':
  assumes eq: "mbuf2t_take f z (mbuf2_add xs ys buf) nts = (z', buf', nts')"
    and bounded: "pred_mapping (λx. x  j) P" "pred_mapping (λx. x  j') P'"
    and rm: "rel_mapping (≤) P P'"
    and pre_buf: "wf_mbuf2' σ P V j n R φ ψ buf"
    and pre_nts: "list_all2 (λi t. t = τ σ i) [min (progress σ P φ j) (progress σ P ψ j)..<j'] nts"
    and xs: "list_all2 (λi. qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) i φ))
      [progress σ P φ j..<progress σ P' φ j'] xs"
    and ys: "list_all2 (λi. qtable n (Formula.fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) i ψ))
      [progress σ P ψ j..<progress σ P' ψ j'] ys"
    and "j  j'"
  shows "wf_mbuf2' σ P' V j' n R φ ψ buf'"
    and "wf_ts σ P' j' φ ψ nts'"
  using pre_buf pre_nts bounded rm unfolding wf_mbuf2'_def wf_ts_def
  by (auto intro!: mbuf2t_take_eqD[OF eq] wf_mbuf2_add[OF _ xs ys] progress_mono_gen[OF j  j' rm]
      progress_le_gen)

lemma ok_0_atms: "ok 0 mr  regex.atms (from_mregex mr []) = {}"
  by (induct mr) auto

lemma ok_0_progress: "ok 0 mr  progress_regex σ P (from_mregex mr []) j = j"
  by (drule ok_0_atms) (auto simp: progress_regex_def)

lemma atms_empty_atms: "safe_regex m g r  atms r = {}  regex.atms r = {}"
  by (induct r rule: safe_regex_induct) (auto split: if_splits simp: cases_Neg_iff)

lemma atms_empty_progress: "safe_regex m g r  atms r = {}  progress_regex σ P r j = j"
  by (drule atms_empty_atms) (auto simp: progress_regex_def)

lemma to_mregex_empty_progress: "safe_regex m g r  to_mregex r = (mr, []) 
  progress_regex σ P r j = j"
  using from_mregex_eq ok_0_progress to_mregex_ok atms_empty_atms by fastforce

lemma progress_regex_le: "pred_mapping (λx. x  j) P  progress_regex σ P r j  j"
  by (auto intro!: progress_le_gen simp: Min_le_iff progress_regex_def)

lemma Neg_acyclic: "formula.Neg x = x  P"
  by (induct x) auto

lemma case_Neg_in_iff: "x  (case y of formula.Neg φ'  {φ'} | _  {})  y = formula.Neg x"
  by (cases y) auto

lemma atms_nonempty_progress:
  "safe_regex m g r  atms r  {}  (λφ. progress σ P φ j) ` atms r = (λφ. progress σ P φ j) ` regex.atms r"
  by (frule atms_empty_atms; simp)
    (auto 0 3 simp: atms_def image_iff case_Neg_in_iff elim!: disjE_Not2 dest: safe_regex_safe_formula)

lemma to_mregex_nonempty_progress: "safe_regex m g r  to_mregex r = (mr, φs)  φs  [] 
  progress_regex σ P r j = (MIN φset φs. progress σ P φ j)"
  using atms_nonempty_progress to_mregex_ok unfolding progress_regex_def by fastforce

lemma to_mregex_progress: "safe_regex m g r  to_mregex r = (mr, φs) 
  progress_regex σ P r j = (if φs = [] then j else (MIN φset φs. progress σ P φ j))"
  using to_mregex_nonempty_progress to_mregex_empty_progress unfolding progress_regex_def by auto

lemma mbufnt_take_add':
  assumes eq: "mbufnt_take f z (mbufn_add xss buf) nts = (z', buf', nts')"
    and bounded: "pred_mapping (λx. x  j) P" "pred_mapping (λx. x  j') P'"
    and rm: "rel_mapping (≤) P P'"
    and safe: "safe_regex m g r"
    and mr: "to_mregex r = (mr, φs)"
    and pre_buf: "wf_mbufn' σ P V j n R r buf"
    and pre_nts: "list_all2 (λi t. t = τ σ i) [progress_regex σ P r j..<j'] nts"
    and xss: "list_all3 list_all2
     (map (λφ i. qtable n (fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) i φ)) φs)
     (map2 upt (map (λφ. progress σ P φ j) φs) (map (λφ. progress σ P' φ j') φs)) xss"
    and "j  j'"
  shows "wf_mbufn' σ P' V j' n R r buf'"
    and "wf_ts_regex σ P' j' r nts'"
  using pre_buf pre_nts bounded rm mr safe xss j  j'  unfolding wf_mbufn'_def wf_ts_regex_def
  using atms_empty_progress[of m g r] to_mregex_ok[OF mr]
  by (auto 0 3 simp: list.rel_map to_mregex_empty_progress to_mregex_nonempty_progress Mini_def
      intro: progress_mono_gen[OF j  j' rm] list.rel_refl_strong progress_le_gen
      dest: list_all2_lengthD elim!: mbufnt_take_eqD[OF eq wf_mbufn_add])

lemma mbuf2t_take_add_induct'[consumes 6, case_names base step]:
  assumes eq: "mbuf2t_take f z (mbuf2_add xs ys buf) nts = (z', buf', nts')"
    and bounded: "pred_mapping (λx. x  j) P" "pred_mapping (λx. x  j') P'"
    and rm: "rel_mapping (≤) P P'"
    and pre_buf: "wf_mbuf2' σ P V j n R φ ψ buf"
    and pre_nts: "list_all2 (λi t. t = τ σ i) [min (progress σ P φ j) (progress σ P ψ j)..<j'] nts"
    and xs: "list_all2 (λi. qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) i φ))
      [progress σ P φ j..<progress σ P' φ j'] xs"
    and ys: "list_all2 (λi. qtable n (Formula.fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) i ψ))
      [progress σ P ψ j..<progress σ P' ψ j'] ys"
    and "j  j'"
    and base: "U (min (progress σ P φ j) (progress σ P ψ j)) z"
    and step: "k X Y z. min (progress σ P φ j) (progress σ P ψ j)  k 
      Suc k  progress σ P' φ j'  Suc k  progress σ P' ψ j' 
      qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) k φ) X 
      qtable n (Formula.fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) k ψ) Y 
      U k z  U (Suc k) (f X Y (τ σ k) z)"
  shows "U (min (progress σ P' φ j') (progress σ P' ψ j')) z'"
  using pre_buf pre_nts bounded rm unfolding wf_mbuf2'_def
  by (blast intro!: mbuf2t_take_induct[OF eq] wf_mbuf2_add[OF _ xs ys] progress_mono_gen[OF j  j' rm]
      progress_le_gen base step)

lemma mbufnt_take_add_induct'[consumes 6, case_names base step]:
  assumes eq: "mbufnt_take f z (mbufn_add xss buf) nts = (z', buf', nts')"
    and bounded: "pred_mapping (λx. x  j) P" "pred_mapping (λx. x  j') P'"
    and rm: "rel_mapping (≤) P P'"
    and safe: "safe_regex m g r"
    and mr: "to_mregex r = (mr, φs)"
    and pre_buf: "wf_mbufn' σ P V j n R r buf"
    and pre_nts: "list_all2 (λi t. t = τ σ i) [progress_regex σ P r j..<j'] nts"
    and xss: "list_all3 list_all2
     (map (λφ i. qtable n (fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) i φ)) φs)
     (map2 upt (map (λφ. progress σ P φ j) φs) (map (λφ. progress σ P' φ j') φs)) xss"
    and "j  j'"
    and base: "U (progress_regex σ P r j) z"
    and step: "k Xs z. progress_regex σ P r j  k  Suc k  progress_regex σ P' r j' 
      list_all2 (λφ. qtable n (Formula.fv φ) (mem_restr R) (λv. Formula.sat σ V (map the v) k φ)) φs Xs 
      U k z  U (Suc k) (f Xs (τ σ k) z)"
  shows "U (progress_regex σ P' r j') z'"
  using pre_buf pre_nts bounded rm j  j' to_mregex_progress[OF safe mr, of σ P' j'] to_mregex_empty_progress[OF safe, of mr σ P j] base
  unfolding wf_mbufn'_def mr prod.case
  by (fastforce dest!: mbufnt_take_induct[OF eq wf_mbufn_add[OF _ xss] pre_nts, of U]
      simp: list.rel_map le_imp_diff_is_add ac_simps Mini_def
      intro: progress_mono_gen[OF j  j' rm] list.rel_refl_strong progress_le_gen
      intro!: base step  dest: list_all2_lengthD split: if_splits)

lemma progress_Until_le: "progress σ P (Formula.Until φ I ψ) j  min (progress σ P φ j) (progress σ P ψ j)"
  by (cases "right I") (auto simp: trans_le_add1 intro!: cInf_lower)

lemma progress_MatchF_le: "progress σ P (Formula.MatchF I r) j  progress_regex σ P r j"
  by (cases "right I") (auto simp: trans_le_add1 progress_regex_def intro!: cInf_lower)

lemma list_all2_upt_Cons: "P a x  list_all2 P [Suc a..<b] xs  Suc a  b 
  list_all2 P [a..<b] (x # xs)"
  by (simp add: list_all2_Cons2 upt_eq_Cons_conv)

lemma list_all2_upt_append: "list_all2 P [a..<b] xs  list_all2 P [b..<c] ys 
  a  b  b  c  list_all2 P [a..<c] (xs @ ys)"
  by (induction xs arbitrary: a) (auto simp add: list_all2_Cons2 upt_eq_Cons_conv)

lemma list_all3_list_all2_conv: "list_all3 R xs xs ys = list_all2 (λx. R x x) xs ys"
  by (auto simp: list_all3_conv_all_nth list_all2_conv_all_nth)

lemma map_split_map: "map_split f (map g xs) = map_split (f o g) xs"
  by (induct xs) auto

lemma map_split_alt: "map_split f xs = (map (fst o f) xs, map (snd o f) xs)"
  by (induct xs) (auto split: prod.splits)

lemma fv_formula_of_constraint: "fv (formula_of_constraint (t1, p, c, t2)) = fv_trm t1  fv_trm t2"
  by (induction "(t1, p, c, t2)" rule: formula_of_constraint.induct) simp_all

lemma (in maux) wf_mformula_wf_set: "wf_mformula σ j P V n R φ φ'  wf_set n (Formula.fv φ')"
  unfolding wf_set_def
proof (induction rule: wf_mformula.induct)
  case (AndRel P V n R φ φ' ψ' conf)
  then show ?case by (auto simp: fv_formula_of_constraint dest!: subsetD)
next
  case (Ands P V n R l l_pos l_neg l' buf A_pos A_neg)
  from Ands.IH have "φ'set (l_pos @ map remove_neg l_neg). xfv φ'. x < n"
    by (simp add: list_all2_conv_all_nth all_set_conv_all_nth[of "_ @ _"] del: set_append)
  then have "φ'set (l_pos @ l_neg). xfv φ'. x < n"
    by (auto dest: bspec[where x="remove_neg _"])
  then show ?case using Ands.hyps(2) by auto
next
  case (Agg P V b n R φ φ' y f g0 ω)
  then have "Formula.fvi_trm b f  Formula.fvi b φ'"
    by (auto simp: fvi_trm_iff_fv_trm[where b=b] fvi_iff_fv[where b=b])
  with Agg show ?case by (auto 0 3 simp: Un_absorb2 fvi_iff_fv[where b=b])
next
  case (MatchP r P V n R φs mr mrs buf nts I aux)
  then obtain φs' where conv: "to_mregex r = (mr, φs')" by blast
  with MatchP have "φ'set φs'. xfv φ'. x < n"
    by (simp add: list_all2_conv_all_nth all_set_conv_all_nth[of φs'])
  with conv show ?case
    by (simp add: to_mregex_ok[THEN conjunct1] fv_regex_alt[OF safe_regex _ _ r])
next
  case (MatchF r  P V n R φs mr mrs buf nts I aux)
  then obtain φs' where conv: "to_mregex r = (mr, φs')" by blast
  with MatchF have "φ'set φs'. xfv φ'. x < n"
    by (simp add: list_all2_conv_all_nth all_set_conv_all_nth[of φs'])
  with conv show ?case
    by (simp add: to_mregex_ok[THEN conjunct1] fv_regex_alt[OF safe_regex _ _ r])
qed (auto simp: fvi_Suc split: if_splits)

lemma qtable_mmulti_join:
  assumes pos: "list_all3 (λA Qi X. qtable n A P Qi X  wf_set n A) A_pos Q_pos L_pos"
    and neg: "list_all3 (λA Qi X. qtable n A P Qi X  wf_set n A) A_neg Q_neg L_neg"
    and C_eq: "C = (set A_pos)" and L_eq: "L = L_pos @ L_neg"
    and "A_pos  []" and fv_subset: "(set A_neg)  (set A_pos)"
    and restrict_pos: "x. wf_tuple n C x  P x  list_all (λA. P (restrict A x)) A_pos"
    and restrict_neg: "x. wf_tuple n C x  P x  list_all (λA. P (restrict A x)) A_neg"
    and Qs: "x. wf_tuple n C x  P x  Q x 
      list_all2 (λA Qi. Qi (restrict A x)) A_pos Q_pos 
      list_all2 (λA Qi. ¬ Qi (restrict A x)) A_neg Q_neg"
  shows "qtable n C P Q (mmulti_join n A_pos A_neg L)"
proof (rule qtableI)
  from pos have 1: "list_all2 (λA X. table n A X  wf_set n A) A_pos L_pos"
    by (simp add: list_all3_conv_all_nth list_all2_conv_all_nth qtable_def)
  moreover from neg have "list_all2 (λA X. table n A X  wf_set n A) A_neg L_neg"
    by (simp add: list_all3_conv_all_nth list_all2_conv_all_nth qtable_def)
  ultimately have L: "list_all2 (λA X. table n A X  wf_set n A) (A_pos @ A_neg) (L_pos @ L_neg)"
    by (rule list_all2_appendI)
  note in_join_iff = mmulti_join_correct[OF A_pos  [] L]
  from 1 have take_eq: "take (length A_pos) (L_pos @ L_neg) = L_pos"
    by (auto dest!: list_all2_lengthD)
  from 1 have drop_eq: "drop (length A_pos) (L_pos @ L_neg) = L_neg"
    by (auto dest!: list_all2_lengthD)

  note mmulti_join.simps[simp del]
  show "table n C (mmulti_join n A_pos A_neg L)"
    unfolding C_eq L_eq table_def by (simp add: in_join_iff)
  show "Q x" if "x  mmulti_join n A_pos A_neg L" "wf_tuple n C x" "P x" for x
    using that(2,3)
  proof (rule Qs[THEN iffD2, OF _ _ conjI])
    have pos': "list_all2 (λA. (∈) (restrict A x)) A_pos L_pos"
      and neg': "list_all2 (λA. (∉) (restrict A x)) A_neg L_neg"
      using that(1) unfolding L_eq in_join_iff take_eq drop_eq by simp_all
    show "list_all2 (λA Qi. Qi (restrict A x)) A_pos Q_pos"
      using pos pos' restrict_pos that(2,3)
      by (simp add: list_all3_conv_all_nth list_all2_conv_all_nth list_all_length qtable_def)
    have fv_subset': "i. i < length A_neg  A_neg ! i  C"
      using fv_subset unfolding C_eq by (auto simp: Sup_le_iff)
    show "list_all2 (λA Qi. ¬ Qi (restrict A x)) A_neg Q_neg"
      using neg neg' restrict_neg that(2,3)
      by (auto simp: list_all3_conv_all_nth list_all2_conv_all_nth list_all_length qtable_def
          wf_tuple_restrict_simple[OF _ fv_subset'])
  qed
  show "x  mmulti_join n A_pos A_neg L" if "wf_tuple n C x" "P x" "Q x" for x
    unfolding L_eq in_join_iff take_eq drop_eq
  proof (intro conjI)
    from that have pos': "list_all2 (λA Qi. Qi (restrict A x)) A_pos Q_pos"
      and neg': "list_all2 (λA Qi. ¬ Qi (restrict A x)) A_neg Q_neg"
      using Qs[THEN iffD1] by auto
    show "wf_tuple n (Aset A_pos. A) x"
      using wf_tuple n C x unfolding C_eq by simp
    show "list_all2 (λA. (∈) (restrict A x)) A_pos L_pos"
      using pos pos' restrict_pos that(1,2)
      by (simp add: list_all3_conv_all_nth list_all2_conv_all_nth list_all_length qtable_def
          C_eq wf_tuple_restrict_simple[OF _ Sup_upper])
    show "list_all2 (λA. (∉) (restrict A x)) A_neg L_neg"
      using neg neg' restrict_neg that(1,2)
      by (auto simp: list_all3_conv_all_nth list_all2_conv_all_nth list_all_length qtable_def)
  qed
qed

lemma nth_filter: "i < length (filter P xs) 
  (i'. i' < length xs  P (xs ! i')  Q (xs ! i'))  Q (filter P xs ! i)"
  by (metis (lifting) in_set_conv_nth set_filter mem_Collect_eq)

lemma nth_partition: "i < length xs 
  (i'. i' < length (filter P xs)  Q (filter P xs ! i')) 
  (i'. i' < length (filter (Not  P) xs)  Q (filter (Not  P) xs ! i'))  Q (xs ! i)"
  by (metis (no_types, lifting) in_set_conv_nth set_filter mem_Collect_eq comp_apply)

lemma qtable_bin_join:
  assumes "qtable n A P Q1 X" "qtable n B P Q2 Y" "¬ b  B  A" "C = A  B"
    "x. wf_tuple n C x  P x  P (restrict A x)  P (restrict B x)"
    "x. b  wf_tuple n C x  P x  Q x  Q1 (restrict A x)  Q2 (restrict B x)"
    "x. ¬ b  wf_tuple n C x  P x  Q x  Q1 (restrict A x)  ¬ Q2 (restrict B x)"
  shows "qtable n C P Q (bin_join n A X b B Y)"
  using qtable_join[OF assms] bin_join_table[of n A X B Y b] assms(1,2)
  by (auto simp add: qtable_def)

lemma restrict_update: "y  A  y < length x  restrict A (x[y:=z]) = restrict A x"
  unfolding restrict_def by (auto simp add: nth_list_update)

lemma qtable_assign:
  assumes "qtable n A P Q X"
    "y < n" "insert y A = A'" "y  A"
    "x'. wf_tuple n A' x'  P x'  P (restrict A x')"
    "x. wf_tuple n A x  P x  Q x  Q' (x[y:=Some (f x)])"
    "x'. wf_tuple n A' x'  P x'  Q' x'  Q (restrict A x')  x' ! y = Some (f (restrict A x'))"
  shows "qtable n A' P Q' ((λx. x[y:=Some (f x)]) ` X)" (is "qtable _ _ _ _ ?Y")
proof (rule qtableI)
  from assms(1) have "table n A X" unfolding qtable_def by simp
  then show "table n A' ?Y"
    unfolding table_def wf_tuple_def using assms(2,3)
    by (auto simp: nth_list_update)
next
  fix x'
  assume "x'  ?Y" "wf_tuple n A' x'" "P x'"
  then obtain x where "x  X" and x'_eq: "x' = x[y:=Some (f x)]" by blast
  then have "wf_tuple n A x"
    using assms(1) unfolding qtable_def table_def by blast
  then have "y < length x" using assms(2) by (simp add: wf_tuple_def)
  with wf_tuple n A x have "restrict A x' = x"
    unfolding x'_eq by (simp add: restrict_update[OF assms(4)] restrict_idle)
  with wf_tuple n A' x' P x' have "P x"
    using assms(5) by blast
  with wf_tuple n A x x  X have "Q x"
    using assms(1) by (elim in_qtableE)
  with wf_tuple n A x P x show "Q' x'"
    unfolding x'_eq by (rule assms(6))
next
  fix x'
  assume "wf_tuple n A' x'" "P x'" "Q' x'"
  then have "wf_tuple n A (restrict A x')"
    using assms(3) by (auto intro!: wf_tuple_restrict_simple)
  moreover have "P (restrict A x')"
    using wf_tuple n A' x' P x' by (rule assms(5))
  moreover have "Q (restrict A x')" and y: "x' ! y = Some (f (restrict A x'))"
    using wf_tuple n A' x' P x' Q' x' by (auto dest!: assms(7))
  ultimately have "restrict A x'  X" by (intro in_qtableI[OF assms(1)])
  moreover have "x' = (restrict A x')[y:=Some (f (restrict A x'))]"
    using y assms(2,3) wf_tuple n A (restrict A x') wf_tuple n A' x'
    by (auto simp: list_eq_iff_nth_eq wf_tuple_def nth_list_update nth_restrict)
  ultimately show "x'  ?Y" by simp
qed

lemma sat_the_update: "y  fv φ  Formula.sat σ V (map the (x[y:=z])) i φ = Formula.sat σ V (map the x) i φ"
  by (rule sat_fv_cong) (metis map_update nth_list_update_neq)

lemma progress_constraint: "progress σ P (formula_of_constraint c) j = j"
  by (induction rule: formula_of_constraint.induct) simp_all

lemma qtable_filter:
  assumes "qtable n A P Q X"
    "x. wf_tuple n A x  P x  Q x  R x  Q' x"
  shows "qtable n A P Q' (Set.filter R X)" (is "qtable _ _ _ _ ?Y")
proof (rule qtableI)
  from assms(1) have "table n A X"
    unfolding qtable_def by simp
  then show "table n A ?Y"
    unfolding table_def wf_tuple_def by simp
next
  fix x
  assume "x  ?Y" "wf_tuple n A x" "P x"
  with assms show "Q' x" by (auto elim!: in_qtableE)
next
  fix x
  assume "wf_tuple n A x" "P x" "Q' x"
  with assms show "x  Set.filter R X" by (auto intro!: in_qtableI)
qed

lemma eval_constraint_sat_eq: "wf_tuple n A x  fv_trm t1  A  fv_trm t2  A 
  iA. i < n  eval_constraint (t1, p, c, t2) x =
    Formula.sat σ V (map the x) i (formula_of_constraint (t1, p, c, t2))"
  by (induction "(t1, p, c, t2)" rule: formula_of_constraint.induct)
    (simp_all add: meval_trm_eval_trm)

declare progress_le_gen[simp]

definition "wf_envs σ j P P' V db =
  (dom V = dom P 
   Mapping.keys db = dom P  {p. p  fst ` Γ σ j} 
   rel_mapping (≤) P P' 
   pred_mapping (λi. i  j) P 
   pred_mapping (λi. i  Suc j) P' 
   (p  Mapping.keys db - dom P. the (Mapping.lookup db p) = [{ts. (p, ts)  Γ σ j}]) 
   (p  dom P. list_all2 (λi X. X = the (V p) i) [the (P p)..<the (P' p)] (the (Mapping.lookup db p))))"


lift_definition mk_db :: "(Formula.name × event_data list) set  Formula.database" is
  "λX p. (if p  fst ` X then Some [{ts. (p, ts)  X}] else None)" .

lemma wf_envs_mk_db: "wf_envs σ j Map.empty Map.empty Map.empty (mk_db (Γ σ j))"
  unfolding wf_envs_def mk_db_def
  by transfer (force split: if_splits simp: image_iff rel_mapping_alt)

lemma wf_envs_update:
  assumes wf_envs: "wf_envs σ j P P' V db"
    and m_eq: "m = Formula.nfv φ"
    and in_fv: "{0 ..< m}  fv φ"
    and xs: "list_all2 (λi. qtable m (Formula.fv φ) (mem_restr UNIV) (λv. Formula.sat σ V (map the v) i φ))
      [progress σ P φ j..<progress σ P' φ (Suc j)] xs"
  shows "wf_envs σ j (P(p  progress σ P φ j)) (P'(p  progress σ P' φ (Suc j)))
    (V(p  λi. {v. length v = m  Formula.sat σ V v i φ}))
    (Mapping.update p (map (image (map the)) xs) db)"
  unfolding wf_envs_def
proof (intro conjI ballI, goal_cases)
  case 3
  from assms show ?case
    by (auto simp: wf_envs_def pred_mapping_alt progress_le progress_mono_gen
        intro!: rel_mapping_map_upd)
next
  case (6 p')
  with assms show ?case by (cases "p'  dom P") (auto simp: wf_envs_def lookup_update')
next
  case (7 p')
  from xs in_fv have "list_all2 (λx y. map the ` y = {v. length v = m  Formula.sat σ V v x φ})
      [progress σ P φ j..<progress σ P' φ (Suc j)] xs"
    by (elim list.rel_mono_strong) (auto 0 3 simp: wf_tuple_def nth_append
      elim!: in_qtableE in_qtableI intro!: image_eqI[where x="map Some _"])
  moreover have "list_all2 (λi X. X = the (V p') i) [the (P p')..<the (P' p')] (the (Mapping.lookup db p'))"
    if "p  p'"
  proof -
    from that 7 have "p'  dom P" by simp
    with wf_envs show ?thesis by (simp add: wf_envs_def)
  qed
  ultimately show ?case
    by (simp add: list.rel_map image_iff lookup_update')
qed (use assms in auto simp: wf_envs_def)

lemma wf_envs_P_simps[simp]:
   "wf_envs σ j P P' V db  pred_mapping (λi. i  j) P"
   "wf_envs σ j P P' V db  pred_mapping (λi. i  Suc j) P'"
   "wf_envs σ j P P' V db  rel_mapping (≤) P P'"
  unfolding wf_envs_def by auto

lemma wf_envs_progress_le[simp]:
   "wf_envs σ j P P' V db  progress σ P φ j  j"
   "wf_envs σ j P P' V db  progress σ P' φ (Suc j)  Suc j"
  unfolding wf_envs_def by auto

lemma wf_envs_progress_regex_le[simp]:
   "wf_envs σ j P P' V db  progress_regex σ P r j  j"
   "wf_envs σ j P P' V db  progress_regex σ P' r (Suc j)  Suc j"
  unfolding wf_envs_def by (auto simp: progress_regex_le)

lemma wf_envs_progress_mono[simp]:
   "wf_envs σ j P P' V db  a  b  progress σ P φ a  progress σ P' φ b"
  unfolding wf_envs_def
  by (auto simp: progress_mono_gen)

lemma qtable_wf_tuple_cong: "qtable n A P Q X  A = B  (v. wf_tuple n A v  P v  Q v = Q' v)  qtable n B P Q' X"
  unfolding qtable_def table_def by blast

lemma (in maux) meval:
  assumes "wf_mformula σ j P V n R φ φ'" "wf_envs σ j P P' V db"
  shows "case meval n (τ σ j) db φ of (xs, φn)  wf_mformula σ (Suc j) P' V n R φn φ' 
    list_all2 (λi. qtable n (Formula.fv φ') (mem_restr R) (λv. Formula.sat σ V (map the v) i φ'))
    [progress σ P φ' j..<progress σ P' φ' (Suc j)] xs"
  using assms
proof (induction φ arbitrary: db P P' V n R φ')
  case (MRel rel)
  then show ?case
    by (cases rule: wf_mformula.cases)
      (auto simp add: ball_Un intro: wf_mformula.intros table_eq_rel eq_rel_eval_trm
        in_eq_rel qtable_empty qtable_unit_table intro!: qtableI)
next
  case (MPred e ts)
  then show ?case
  proof (cases "e  dom P")
    case True
    with MPred(2) have "e  Mapping.keys db" "e  dom P'" "e  dom V"
      "list_all2 (λi X. X = the (V e) i) [the (P e)..<the (P' e)]
         (the (Mapping.lookup db e))" unfolding wf_envs_def rel_mapping_alt by blast+
    with MPred(1) True show ?thesis
      by (cases rule: wf_mformula.cases)
        (fastforce intro!: wf_mformula.Pred qtableI bexI[where P="λx. _ = tabulate x 0 n", OF refl]
        elim!: list.rel_mono_strong bexI[rotated] dest: ex_match
        simp: list.rel_map table_def match_wf_tuple in_these_eq match_eval_trm image_iff
          list.map_comp keys_dom_lookup)
  next
    note MPred(1)
    moreover
    case False
    moreover
    from False MPred(2) have "e  dom P'" "e  dom V"
      unfolding wf_envs_def rel_mapping_alt by auto
    moreover
    from False MPred(2) have *: "e  fst ` Γ σ j  e  Mapping.keys db"
      unfolding wf_envs_def by auto
    from False MPred(2) have
      "e  Mapping.keys db  Mapping.lookup db e = Some [{ts. (e, ts)  Γ σ j}]"
      unfolding wf_envs_def keys_dom_lookup by (metis Diff_iff domD option.sel)
    with * have "(case Mapping.lookup db e of None  [{}] | Some xs  xs) = [{ts. (e, ts)  Γ σ j}]"
      by (cases "e  fst ` Γ σ j") (auto simp: image_iff keys_dom_lookup split: option.splits)
    ultimately show ?thesis
      by (cases rule: wf_mformula.cases)
        (fastforce intro!: wf_mformula.Pred qtableI bexI[where P="λx. _ = tabulate x 0 n", OF refl]
        elim!: list.rel_mono_strong bexI[rotated] dest: ex_match
        simp: list.rel_map table_def match_wf_tuple in_these_eq match_eval_trm image_iff list.map_comp)
  qed
next
  case (MLet p m φ1 φ2)
  from MLet.prems(1) obtain φ1' φ2' where Let: "φ' = Formula.Let p φ1' φ2'" and
    1: "wf_mformula σ j P V m UNIV φ1 φ1'" and
    fv: "m = Formula.nfv φ1'" "{0..<m}  fv φ1'" and
    2: "wf_mformula σ j (P(p  progress σ P φ1' j))
      (V(p  λi. {v. length v = m  Formula.sat σ V v i φ1'}))
      n R φ2 φ2'"
    by (cases rule: wf_mformula.cases) auto
  obtain xs φ1_new where e1: "meval m (τ σ j) db φ1 = (xs, φ1_new)" and
      wf1: "wf_mformula σ (Suc j) P' V m UNIV φ1_new φ1'" and
      res1: "list_all2 (λi. qtable m (fv φ1') (mem_restr UNIV) (λv. Formula.sat σ V (map the v) i φ1'))
       [progress σ P φ1' j..<progress σ P' φ1' (Suc j)] xs"
    using MLet(1)[OF 1(1) MLet.prems(2)] by (auto simp: eqTrueI[OF mem_restr_UNIV, abs_def])
  from MLet(2)[OF 2 wf_envs_update[OF MLet.prems(2) fv res1]] wf1 e1 fv
  show ?case unfolding Let
    by (auto simp: fun_upd_def intro!: wf_mformula.Let)
next
  case (MAnd A_φ φ pos A_ψ ψ buf)
  from MAnd.prems show ?case
    by (cases rule: wf_mformula.cases)
      (auto simp: sat_the_restrict simp del: bin_join.simps
        dest!: MAnd.IH split: if_splits prod.splits intro!: wf_mformula.And qtable_bin_join
        elim: mbuf2_take_add'(1) list.rel_mono_strong[OF mbuf2_take_add'(2)])
next
  case (MAndAssign φ conf)
  from MAndAssign.prems obtain φ'' x t ψ'' where
    wf_envs: "wf_envs σ j P P' V db" and
    φ'_eq: "φ' = formula.And φ'' ψ''" and
    wf_φ: "wf_mformula σ j P V n R φ φ''" and
    "x < n" and
    "x  fv φ''" and
    fv_t_subset: "fv_trm t  fv φ''" and
    conf: "(x, t) = conf" and
    ψ''_eqs: "ψ'' = formula.Eq (trm.Var x) t  ψ'' = formula.Eq t (trm.Var x)"
    by (cases rule: wf_mformula.cases)
  from wf_φ wf_envs obtain xs φn where
    meval_eq: "meval n (τ σ j) db φ = (xs, φn)" and
    wf_φn: "wf_mformula σ (Suc j) P' V n R φn φ''" and
    xs: "list_all2 (λi. qtable n (fv φ'') (mem_restr R) (λv. Formula.sat σ V (map the v) i φ''))
        [progress σ P φ'' j..<progress σ P' φ'' (Suc j)] xs"
    by (auto dest!: MAndAssign.IH)
  have progress_eqs:
      "progress σ P φ' j = progress σ P φ'' j"
      "progress σ P' φ' (Suc j) = progress σ P' φ'' (Suc j)"
    using ψ''_eqs wf_envs_progress_le[OF wf_envs] by (auto simp: φ'_eq)

  show ?case proof (simp add: meval_eq, intro conjI)
    show "wf_mformula σ (Suc j) P' V n R (MAndAssign φn conf) φ'"
      unfolding φ'_eq
      by (rule wf_mformula.AndAssign) fact+
  next
    show "list_all2 (λi. qtable n (fv φ') (mem_restr R) (λv. Formula.sat σ V (map the v) i φ'))
        [progress σ P φ' j..<progress σ P' φ' (Suc j)] (map ((`) (eval_assignment conf)) xs)"
      unfolding list.rel_map progress_eqs conf[symmetric] eval_assignment.simps
      using xs
    proof (rule list.rel_mono_strong)
      fix i X
      assume qtable: "qtable n (fv φ'') (mem_restr R) (λv. Formula.sat σ V (map the v) i φ'') X"
      then show "qtable n (fv φ') (mem_restr R) (λv. Formula.sat σ V (map the v) i φ')
          ((λy. y[x := Some (meval_trm t y)]) ` X)"
      proof (rule qtable_assign)
        show "x < n" by fact
        show "insert x (fv φ'') = fv φ'"
          using ψ''_eqs fv_t_subset by (auto simp: φ'_eq)
        show "x  fv φ''" by fact
      next
        fix v
        assume wf_v: "wf_tuple n (fv φ') v" and "mem_restr R v"
        then show "mem_restr R (restrict (fv φ'') v)" by simp

        assume sat: "Formula.sat σ V (map the v) i φ'"
        then have A: "Formula.sat σ V (map the (restrict (fv φ'') v)) i φ''" (is ?A)
          by (simp add: φ'_eq sat_the_restrict)
        have "map the v ! x = Formula.eval_trm (map the v) t"
          using sat ψ''_eqs by (auto simp: φ'_eq)
        also have "... = Formula.eval_trm (map the (restrict (fv φ'') v)) t"
          using fv_t_subset by (auto simp: map_the_restrict intro!: eval_trm_fv_cong)
        finally have "map the v ! x = meval_trm t (restrict (fv φ'') v)"
          using meval_trm_eval_trm[of n "fv φ''" "restrict (fv φ'') v" t]
            fv_t_subset wf_v wf_mformula_wf_set[unfolded wf_set_def, OF wf_φ]
          by (fastforce simp: φ'_eq intro!: wf_tuple_restrict)
        then have B: "v ! x = Some (meval_trm t (restrict (fv φ'') v))" (is ?B)
          using ψ''_eqs wf_v x < n by (auto simp: wf_tuple_def φ'_eq)
        from A B show "?A  ?B" ..
      next
        fix v
        assume wf_v: "wf_tuple n (fv φ'') v" and "mem_restr R v"
          and sat: "Formula.sat σ V (map the v) i φ''"
        let ?v = "v[x := Some (meval_trm t v)]"
        from sat have A: "Formula.sat σ V (map the ?v) i φ''"
          using x  fv φ'' by (simp add: sat_the_update)
        have "y  fv_trm t  x  y" for y
          using fv_t_subset x  fv φ'' by auto
        then have B: "Formula.sat σ V (map the ?v) i ψ''"
          using ψ''_eqs meval_trm_eval_trm[of n "fv φ''" v t] x < n
            fv_t_subset wf_v wf_mformula_wf_set[unfolded wf_set_def, OF wf_φ]
          by (auto simp: wf_tuple_def map_update intro!: eval_trm_fv_cong)
        from A B show "Formula.sat σ V (map the ?v) i φ'"
          by (simp add: φ'_eq)
      qed
    qed
  qed
next
  case (MAndRel φ conf)
  from MAndRel.prems show ?case
    by (cases rule: wf_mformula.cases)
      (auto simp: progress_constraint progress_le list.rel_map fv_formula_of_constraint
        Un_absorb2 wf_mformula_wf_set[unfolded wf_set_def] split: prod.splits
        dest!: MAndRel.IH[where db=db and P=P and P'=P'] eval_constraint_sat_eq[THEN iffD2]
        intro!: wf_mformula.AndRel
        elim!: list.rel_mono_strong qtable_filter eval_constraint_sat_eq[THEN iffD1])
next
  case (MAnds A_pos A_neg l buf)
  note mbufn_take.simps[simp del] mbufn_add.simps[simp del] mmulti_join.simps[simp del]

  from MAnds.prems obtain pos neg l' where
    wf_l: "list_all2 (wf_mformula σ j P V n R) l (pos @ map remove_neg neg)" and
    wf_buf: "wf_mbufn (progress σ P (formula.Ands l') j) (map (λψ. progress σ P ψ j) (pos @ map remove_neg neg))
      (map (λψ i. qtable n (fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) i ψ)) (pos @ map remove_neg neg)) buf" and
    posneg: "(pos, neg) = partition safe_formula l'" and
    "pos  []" and
    safe_neg: "list_all safe_formula (map remove_neg neg)" and
    A_eq: "A_pos = map fv pos" "A_neg = map fv neg" and
    fv_subset: " (set A_neg)   (set A_pos)" and
    "φ' = Formula.Ands l'"
    by (cases rule: wf_mformula.cases) simp
  have progress_eq: "progress σ P' (formula.Ands l') (Suc j) =
      Mini (progress σ P (formula.Ands l') j) (map (λψ. progress σ P' ψ (Suc j)) (pos @ map remove_neg neg))"
    using pos  [] posneg
    by (auto simp: Mini_def image_Un[symmetric] Collect_disj_eq[symmetric] intro!: arg_cong[where f=Min])

  have join_ok: "qtable n ( (fv ` set l')) (mem_restr R)
        (λv. list_all (Formula.sat σ V (map the v) k) l')
        (mmulti_join n A_pos A_neg L)"
    if args_ok: "list_all2 (λx. qtable n (fv x) (mem_restr R) (λv. Formula.sat σ V (map the v) k x))
        (pos @ map remove_neg neg) L"
    for k L
  proof (rule qtable_mmulti_join)
    let ?ok = "λA Qi X. qtable n A (mem_restr R) Qi X  wf_set n A"
    let ?L_pos = "take (length A_pos) L"
    let ?L_neg = "drop (length A_pos) L"
    have 1: "length pos  length L"
      using args_ok by (auto dest!: list_all2_lengthD)
    show "list_all3 ?ok A_pos (map (λψ v. Formula.sat σ V (map the v) k ψ) pos) ?L_pos"
      using args_ok wf_l unfolding A_eq
      by (auto simp add: list_all3_conv_all_nth list_all2_conv_all_nth nth_append
          split: if_splits intro!: wf_mformula_wf_set[of σ j P V n R]
          dest: order.strict_trans2[OF _ 1])
    from args_ok have prems_neg: "list_all2 (λψ. qtable n (fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) k (remove_neg ψ))) neg ?L_neg"
      by (auto simp: A_eq list_all2_append1 list.rel_map)
    show "list_all3 ?ok A_neg (map (λψ v. Formula.sat σ V (map the v) k (remove_neg ψ)) neg) ?L_neg"
      using prems_neg wf_l unfolding A_eq
      by (auto simp add: list_all3_conv_all_nth list_all2_conv_all_nth list_all_length nth_append less_diff_conv
          split: if_splits intro!: wf_mformula_wf_set[of σ j P V n R _ "remove_neg _", simplified])
    show "(fv ` set l') = (set A_pos)"
      using fv_subset posneg unfolding A_eq by auto
    show "L = take (length A_pos) L @ drop (length A_pos) L" by simp
    show "A_pos  []" using pos  [] A_eq by simp

    fix x :: "event_data tuple"
    assume "wf_tuple n ( (fv ` set l')) x" and "mem_restr R x"
    then show "list_all (λA. mem_restr R (restrict A x)) A_pos"
      and "list_all (λA. mem_restr R (restrict A x)) A_neg"
      by (simp_all add: list.pred_set)

    have "list_all Formula.is_Neg neg"
      using posneg safe_neg
      by (auto 0 3 simp add: list.pred_map elim!: list.pred_mono_strong
          intro: formula.exhaust[of ψ "Formula.is_Neg ψ" for ψ])
    then have "list_all (λψ. Formula.sat σ V (map the v) i (remove_neg ψ) 
      ¬ Formula.sat σ V (map the v) i ψ) neg" for v i
      by (fastforce simp: Formula.is_Neg_def elim!: list.pred_mono_strong)
    then show "list_all (Formula.sat σ V (map the x) k) l' =
       (list_all2 (λA Qi. Qi (restrict A x)) A_pos
         (map (λψ v. Formula.sat σ V (map the v) k ψ) pos) 
        list_all2 (λA Qi. ¬ Qi (restrict A x)) A_neg
         (map (λψ v. Formula.sat σ V (map the v) k
                       (remove_neg ψ))
           neg))"
      using posneg
      by (auto simp add: A_eq list_all2_conv_all_nth list_all_length sat_the_restrict
          elim: nth_filter nth_partition[where P=safe_formula and Q="Formula.sat _ _ _ _"])
  qed fact

  from MAnds.prems(2) show ?case
    unfolding φ' = Formula.Ands l'
    by (auto 0 3 simp add: list.rel_map progress_eq map2_map_map list_all3_map
        list_all3_list_all2_conv list.pred_map
        simp del: set_append map_append progress_simps split: prod.splits
        intro!: wf_mformula.Ands[OF _ _ posneg pos  [] safe_neg A_eq fv_subset]
        list.rel_mono_strong[OF wf_l] wf_mbufn_add[OF wf_buf]
        list.rel_flip[THEN iffD1, OF list.rel_mono_strong, OF wf_l]
        list.rel_refl join_ok[unfolded list.pred_set]
        dest!: MAnds.IH[OF _ _ MAnds.prems(2), rotated]
        elim!: wf_mbufn_take list_all2_appendI
        elim: mbufn_take_induct[OF _ wf_mbufn_add[OF wf_buf]])
next
  case (MOr φ ψ buf)
  from MOr.prems show ?case
    by (cases rule: wf_mformula.cases)
      (auto dest!: MOr.IH split: if_splits prod.splits intro!: wf_mformula.Or qtable_union
        elim: mbuf2_take_add'(1) list.rel_mono_strong[OF mbuf2_take_add'(2)])
next
  case (MNeg φ)
  have *: "qtable n {} (mem_restr R) (λv. P v) X 
    ¬ qtable n {} (mem_restr R) (λv. ¬ P v) empty_table  x  X  False" for P x X
    using nullary_qtable_cases qtable_unit_empty_table by fastforce
  from MNeg.prems show ?case
    by (cases rule: wf_mformula.cases)
      (auto 0 4 intro!: wf_mformula.Neg dest!: MNeg.IH
        simp add: list.rel_map
        dest: nullary_qtable_cases qtable_unit_empty_table intro!: qtable_empty_unit_table
        elim!: list.rel_mono_strong elim: *)
next
  case (MExists φ)
  from MExists.prems show ?case
    by (cases rule: wf_mformula.cases)
      (force simp: list.rel_map fvi_Suc sat_fv_cong nth_Cons'
        intro!: wf_mformula.Exists dest!: MExists.IH qtable_project_fv
        elim!: list.rel_mono_strong table_fvi_tl qtable_cong sat_fv_cong[THEN iffD1, rotated -1]
        split: if_splits)+
next
  case (MAgg g0 y ω b f φ)
  from MAgg.prems show ?case
    using wf_mformula_wf_set[OF MAgg.prems(1), unfolded wf_set_def]
    by (cases rule: wf_mformula.cases)
      (auto 0 3 simp add: list.rel_map simp del: sat.simps fvi.simps split: prod.split
        intro!: wf_mformula.Agg qtable_eval_agg dest!: MAgg.IH elim!: list.rel_mono_strong)
next
  case (MPrev I φ first buf nts)
  from MPrev.prems show ?case
  proof (cases rule: wf_mformula.cases)
    case (Prev ψ)
    let ?xs = "fst (meval n (τ σ j) db φ)"
    let  = "snd (meval n (τ σ j) db φ)"
    let ?ls = "fst (mprev_next I (buf @ ?xs) (nts @ [τ σ j]))"
    let ?rs = "fst (snd (mprev_next I (buf @ ?xs) (nts @ [τ σ j])))"
    let ?ts = "snd (snd (mprev_next I (buf @ ?xs) (nts @ [τ σ j])))"
    let ?P = "λi X. qtable n (fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) i ψ) X"
    let ?min = "min (progress σ P' ψ (Suc j)) (Suc j - 1)"
    from Prev MPrev.IH[OF _ MPrev.prems(2), of n R ψ] have IH: "wf_mformula σ (Suc j) P' V n R  ψ" and
      "list_all2 ?P [progress σ P ψ j..<progress σ P' ψ (Suc j)] ?xs" by auto
    with Prev(4,5) MPrev.prems(2) have "list_all2 (λi X. if mem (τ σ (Suc i) - τ σ i) I then ?P i X else X = empty_table)
        [min (progress σ P ψ j) (j - 1)..<?min] ?ls 
       list_all2 ?P [?min..<progress σ P' ψ (Suc j)] ?rs 
       list_all2 (λi t. t = τ σ i) [?min..<Suc j] ?ts"
      by (intro mprev) (auto intro!: list_all2_upt_append list_all2_appendI order.trans[OF min.cobounded1])
    moreover have "min (Suc (progress σ P ψ j)) j = Suc (min (progress σ P ψ j) (j-1))" if "j > 0"
      using that by auto
    ultimately show ?thesis using Prev(1,3) MPrev.prems(2) IH
      by (auto simp: map_Suc_upt[symmetric] upt_Suc[of 0] list.rel_map qtable_empty_iff
          simp del: upt_Suc elim!: wf_mformula.Prev list.rel_mono_strong
          split: prod.split if_split_asm)
  qed
next
  case (MNext I φ first nts)
  from MNext.prems show ?case
  proof (cases rule: wf_mformula.cases)
    case (Next ψ)

    have min[simp]:
      "min (progress σ P ψ j - Suc 0) (j - Suc 0) = progress σ P ψ j - Suc 0"
      "min (progress σ P' ψ (Suc j) - Suc 0) j = progress σ P' ψ (Suc j) - Suc 0"
      using wf_envs_progress_le[OF MNext.prems(2), of ψ] by auto

    let ?xs = "fst (meval n (τ σ j) db φ)"
    let ?ys = "case (?xs, first) of (_ # xs, True)  xs | _  ?xs"
    let  = "snd (meval n (τ σ j) db φ)"
    let ?ls = "fst (mprev_next I ?ys (nts @ [τ σ j]))"
    let ?rs = "fst (snd (mprev_next I ?ys (nts @ [τ σ j])))"
    let ?ts = "snd (snd (mprev_next I ?ys (nts @ [τ σ j])))"
    let ?P = "λi X. qtable n (fv ψ) (mem_restr R) (λv. Formula.sat σ V (map the v) i ψ) X"
    let ?min = "min (progress σ P' ψ (Suc j) - 1) (Suc j - 1)"
    from Next MNext.IH[OF _ MNext.prems(2), of n R ψ] have IH: "wf_mformula σ (Suc j) P' V  n R  ψ"
      "list_all2 ?P [progress σ P ψ j..<progress σ P' ψ (Suc j)] ?xs" by auto
    with Next have "list_all2 (λi X. if mem (τ σ (Suc i) - τ σ i) I then ?P (Suc i) X else X = empty_table)
        [progress σ P ψ j - 1..<?min] ?ls 
       list_all2 ?P [Suc ?min..<progress σ P' ψ (Suc j)] ?rs 
       list_all2 (λi t. t = τ σ i) [?min..<Suc j] ?ts" if "progress σ P ψ j < progress σ P' ψ (Suc j)"
      using that wf_envs_progress_le[OF MNext.prems(2), of ψ]
      by (intro mnext) (auto simp: list_all2_Cons2 upt_eq_Cons_conv
          intro!: list_all2_upt_append list_all2_appendI split: list.splits)
    then show ?thesis using wf_envs_progress_le[OF MNext.prems(2), of ψ]
      wf_envs_progress_mono[OF MNext.prems(2), of j "Suc j" ψ, simplified] Next IH
      by (cases "progress σ P' ψ (Suc j) > progress σ P ψ j")
        (auto 0 3 simp: qtable_empty_iff le_Suc_eq le_diff_conv
          elim!: wf_mformula.Next list.rel_mono_strong list_all2_appendI
          split: prod.split list.splits if_split_asm)  (* slow 5 sec*)
  qed
next
  case (MSince args φ ψ buf nts aux)
  note sat.simps[simp del]
  from MSince.prems obtain φ'' φ''' ψ'' I where Since_eq: "φ' = Formula.Since φ''' I ψ''"
    and pos: "if args_pos args then φ''' = φ'' else φ''' = Formula.Neg φ''"
    and pos_eq: "safe_formula φ''' = args_pos args"
    and φ: "wf_mformula σ j P V n R φ φ''"
    and ψ: "wf_mformula σ j P V n R ψ ψ''"
    and fvi_subset: "Formula.fv φ''  Formula.fv ψ''"
    and buf: "wf_mbuf2' σ P V j n R φ'' ψ'' buf"
    and nts: "wf_ts σ P j φ'' ψ'' nts"
    and aux: "wf_since_aux σ V R args φ'' ψ'' aux (progress σ P (Formula.Since φ''' I ψ'') j)"
    and args_ivl: "args_ivl args = I"
    and args_n: "args_n args = n"
    and args_L: "args_L args = Formula.fv φ''"
    and args_R: "args_R args = Formula.fv ψ''"
    by (cases rule: wf_mformula.cases) (auto)
  have φ''': "Formula.fv φ''' = Formula.fv φ''" "progress σ P φ''' j = progress σ P φ'' j"
    "progress σ P' φ''' j = progress σ P' φ'' j" for j
    using pos by (simp_all split: if_splits)
  from MSince.prems(2) have nts_snoc: "list_all2 (λi t. t = τ σ i)
    [min (progress σ P φ'' j) (progress σ P ψ'' j)..<Suc j] (nts @ [τ σ j])"
    using nts unfolding wf_ts_def
    by (auto simp add: wf_envs_progress_le[THEN min.coboundedI1] intro: list_all2_appendI)
  have update: "wf_since_aux σ V R args φ'' ψ'' (snd (zs, aux')) (progress σ P' (Formula.Since φ''' I ψ'') (Suc j)) 
    list_all2 (λi. qtable n (Formula.fv φ'''  Formula.fv ψ'') (mem_restr R)
      (λv. Formula.sat σ V (map the v) i (Formula.Since φ''' I ψ'')))
      [progress σ P (Formula.Since φ''' I ψ'') j..<progress σ P' (Formula.Since φ''' I ψ'') (Suc j)] (fst (zs, aux'))"
    if eval_φ: "fst (meval n (τ σ j) db φ) = xs"
      and eval_ψ: "fst (meval n (τ σ j) db ψ) = ys"
      and eq: "mbuf2t_take (λr1 r2 t (zs, aux).
        case update_since args r1 r2 t aux of (z, x)  (zs @ [z], x))
        ([], aux) (mbuf2_add xs ys buf) (nts @ [τ σ j]) = ((zs, aux'), buf', nts')"
    for xs ys zs aux' buf' nts'
    unfolding progress_simps φ'''
  proof (rule mbuf2t_take_add_induct'[where j=j and j'="Suc j" and z'="(zs, aux')",
      OF eq wf_envs_P_simps[OF MSince.prems(2)] buf nts_snoc],
      goal_cases xs ys _ base step)
    case xs
    then show ?case
      using MSince.IH(1)[OF φ MSince.prems(2)] eval_φ by auto
  next
    case ys
    then show ?case
      using MSince.IH(2)[OF ψ MSince.prems(2)] eval_ψ by auto
  next
    case base
    then show ?case
      using aux by (simp add: φ''')
  next
    case (step k X Y z)
    then show ?case
      using fvi_subset pos
      by (auto 0 3 simp: args_ivl args_n args_L args_R Un_absorb1
          elim!: update_since(1) list_all2_appendI dest!: update_since(2)
          split: prod.split if_splits)
  qed simp
  with MSince.IH(1)[OF φ MSince.prems(2)] MSince.IH(2)[OF ψ MSince.prems(2)] show ?case
    by (auto 0 3 simp: Since_eq split: prod.split
        intro: wf_mformula.Since[OF _ _ pos pos_eq args_ivl args_n args_L args_R fvi_subset]
        elim: mbuf2t_take_add'(1)[OF _ wf_envs_P_simps[OF MSince.prems(2)] buf nts_snoc]
              mbuf2t_take_add'(2)[OF _ wf_envs_P_simps[OF MSince.prems(2)] buf nts_snoc])
next
  case (MUntil args φ ψ buf nts aux)
  note sat.simps[simp del] progress_simps[simp del]
  from MUntil.prems obtain φ'' φ''' ψ'' I where Until_eq: "φ' = Formula.Until φ''' I ψ''"
    and pos: "if args_pos args then φ''' = φ'' else φ''' = Formula.Neg φ''"
    and pos_eq: "safe_formula φ''' = args_pos args"
    and φ: "wf_mformula σ j P V n R φ φ''"
    and ψ: "wf_mformula σ j P V n R ψ ψ''"
    and fvi_subset: "Formula.fv φ''  Formula.fv ψ''"
    and buf: "wf_mbuf2' σ P V j n R φ'' ψ'' buf"
    and nts: "wf_ts σ P j φ'' ψ'' nts"
    and aux: "wf_until_aux σ V R args φ'' ψ'' aux (progress σ P (Formula.Until φ''' I ψ'') j)"
    and args_ivl: "args_ivl args = I"
    and args_n: "args_n args = n"
    and args_L: "args_L args = Formula.fv φ''"
    and args_R: "args_R args = Formula.fv ψ''"
    and length_aux: "progress σ P (Formula.Until φ''' I ψ'') j + length_muaux args aux =
      min (progress σ P φ'' j) (progress σ P ψ'' j)"
    by (cases rule: wf_mformula.cases) (auto)
  define pos where args_pos: "pos = args_pos args"
  have φ''': "progress σ P φ''' j = progress σ P φ'' j"  "progress σ P' φ''' j = progress σ P' φ'' j" for j
    using pos by (simp_all add: progress.simps split: if_splits)
  from MUntil.prems(2) have nts_snoc: "list_all2 (λi t. t = τ σ i)
    [min (progress σ P φ'' j) (progress σ P ψ'' j)..<Suc j] (nts @ [τ σ j])"
    using nts unfolding wf_ts_def
    by (auto simp add: wf_envs_progress_le[THEN min.coboundedI1] intro: list_all2_appendI)
  {
    fix xs ys zs aux' aux'' buf' nts'
    assume eval_φ: "fst (meval n (τ σ j) db φ) = xs"
      and eval_ψ: "fst (meval n (τ σ j) db ψ) = ys"
      and eq1: "mbuf2t_take (add_new_muaux args) aux (mbuf2_add xs ys buf) (nts @ [τ σ j]) =
        (aux', buf', nts')"
      and eq2: "eval_muaux args (case nts' of []  τ σ j | nt # _  nt) aux' = (zs, aux'')"
    define ne where "ne  progress σ P (Formula.Until φ''' I ψ'') j"
    have update1: "wf_until_aux σ V R args φ'' ψ'' aux' (progress σ P (Formula.Until φ''' I ψ'') j) 
      ne + length_muaux args aux' = min (progress σ P' φ''' (Suc j)) (progress σ P' ψ'' (Suc j))"
      using MUntil.IH(1)[OF φ MUntil.prems(2)] eval_φ MUntil.IH(2)[OF ψ MUntil.prems(2)]
        eval_ψ nts_snoc nts_snoc length_aux aux fvi_subset
      unfolding φ'''
      by (elim mbuf2t_take_add_induct'[where j'="Suc j", OF eq1 wf_envs_P_simps[OF MUntil.prems(2)] buf])
        (auto simp: args_n args_L args_R ne_def wf_update_until)
    then obtain cur auxlist' where valid_aux': "valid_muaux args cur aux' auxlist'" and
      cur: "cur = (if ne + length auxlist' = 0 then 0 else τ σ (ne + length auxlist' - 1))" and
      wf_auxlist': "wf_until_auxlist σ V n R pos φ'' I ψ'' auxlist' (progress σ P (Formula.Until φ''' I ψ'') j)"
      unfolding wf_until_aux_def ne_def args_ivl args_n args_pos by auto
    have length_aux': "length_muaux args aux' = length auxlist'"
      using valid_length_muaux[OF valid_aux'] .
    have nts': "wf_ts σ P' (Suc j) φ'' ψ'' nts'"
      using MUntil.IH(1)[OF φ MUntil.prems(2)] eval_φ MUntil.IH(2)[OF ψ MUntil.prems(2)]
        MUntil.prems(2) eval_ψ nts_snoc
      unfolding wf_ts_def
      by (intro mbuf2t_take_eqD(2)[OF eq1]) (auto intro: wf_mbuf2_add buf[unfolded wf_mbuf2'_def])
    define zs'' where "zs'' = fst (eval_until I (case nts' of []  τ σ j | nt # x  nt) auxlist')"
    define auxlist'' where "auxlist'' = snd (eval_until I (case nts' of []  τ σ j | nt # x  nt) auxlist')"
    have current_w_le: "cur  (case nts' of []  τ σ j | nt # x  nt)"
    proof (cases nts')
      case Nil
      have p_le: "min (progress σ P' φ''' (Suc j)) (progress σ P' ψ'' (Suc j)) - 1  j"
        using wf_envs_progress_le[OF MUntil.prems(2)]
        by (auto simp: min_def le_diff_conv)
      then show ?thesis
        unfolding cur conjunct2[OF update1, unfolded length_aux']
        using Nil by auto
    next
      case (Cons nt x)
      have progress_φ''': "progress σ P' φ'' (Suc j) = progress σ P' φ''' (Suc j)"
        using pos by (auto simp add: progress.simps split: if_splits)
      have "nt = τ σ (min (progress σ P' φ'' (Suc j)) (progress σ P' ψ'' (Suc j)))"
        using nts'[unfolded wf_ts_def Cons]
        unfolding list_all2_Cons2 upt_eq_Cons_conv by auto
      then show ?thesis
        unfolding cur conjunct2[OF update1, unfolded length_aux'] Cons progress_φ'''
        by (auto split: if_splits list.splits intro!: τ_mono)
    qed
    have valid_aux'': "valid_muaux args cur aux'' auxlist''"
      using valid_eval_muaux[OF valid_aux' current_w_le eq2, of zs'' auxlist'']
      by (auto simp add: args_ivl zs''_def auxlist''_def)
    have length_aux'': "length_muaux args aux'' = length auxlist''"
      using valid_length_muaux[OF valid_aux''] .
    have eq2': "eval_until I (case nts' of []  τ σ j | nt # _  nt) auxlist' = (zs, auxlist'')"
      using valid_eval_muaux[OF valid_aux' current_w_le eq2, of zs'' auxlist'']
      by (auto simp add: args_ivl zs''_def auxlist''_def)
    have length_aux'_aux'': "length_muaux args aux' = length zs + length_muaux args aux''"
      using eval_until_length[OF eq2'] unfolding length_aux' length_aux'' .
    have "i  progress σ P' (Formula.Until φ''' I ψ'') (Suc j) 
      wf_until_auxlist σ V n R pos φ'' I ψ'' auxlist' i 
      i + length auxlist' = min (progress σ P' φ''' (Suc j)) (progress σ P' ψ'' (Suc j)) 
      wf_until_auxlist σ V n R pos φ'' I ψ'' auxlist'' (progress σ P' (Formula.Until φ''' I ψ'') (Suc j)) 
        i + length zs = progress σ P' (Formula.Until φ''' I ψ'') (Suc j) 
        i + length zs + length auxlist'' = min (progress σ P' φ''' (Suc j)) (progress σ P' ψ'' (Suc j)) 
        list_all2 (λi. qtable n (Formula.fv ψ'') (mem_restr R)
          (λv. Formula.sat σ V (map the v) i (Formula.Until (if pos then φ'' else Formula.Neg φ'') I ψ'')))
          [i..<i + length zs] zs" for i
      using eq2'
    proof (induction auxlist' arbitrary: zs auxlist'' i)
      case Nil
      then show ?case
        by (auto dest!: antisym[OF progress_Until_le])
    next
      case (Cons a aux')
      obtain t a1 a2 where "a = (t, a1, a2)" by (cases a)
      from Cons.prems(2) have aux': "wf_until_auxlist σ V n R pos φ'' I ψ'' aux' (Suc i)"
        by (rule wf_until_aux_Cons)
      from Cons.prems(2) have 1: "t = τ σ i"
        unfolding a = (t, a1, a2) by (rule wf_until_aux_Cons1)
      from Cons.prems(2) have 3: "qtable n (Formula.fv ψ'') (mem_restr R) (λv.
        (ji. j < Suc (i + length aux')  mem (τ σ j - τ σ i) I  Formula.sat σ V (map the v) j ψ'' 
        (k{i..<j}. if pos then Formula.sat σ V (map the v) k φ'' else ¬ Formula.sat σ V (map the v) k φ''))) a2"
        unfolding a = (t, a1, a2) by (rule wf_until_aux_Cons3)
      from Cons.prems(3) have Suc_i_aux': "Suc i + length aux' =
          min (progress σ P' φ''' (Suc j)) (progress σ P' ψ'' (Suc j))"
        by simp
      have "i  progress σ P' (Formula.Until φ''' I ψ'') (Suc j)"
        if "enat (case nts' of []  τ σ j | nt # x  nt)  enat t + right I"
        using that nts' unfolding wf_ts_def progress.simps
        by (auto simp add: 1 list_all2_Cons2 upt_eq_Cons_conv φ'''
            intro!: cInf_lower τ_mono elim!: order.trans[rotated] simp del: upt_Suc split: if_splits list.splits)
      moreover
      have "Suc i  progress σ P' (Formula.Until φ''' I ψ'') (Suc j)"
        if "enat t + right I < enat (case nts' of []  τ σ j | nt # x  nt)"
      proof -
        from that obtain m where m: "right I = enat m" by (cases "right I") auto
        have τ_min:  "τ σ (min j k) = min (τ σ j) (τ σ k)" for k
          by (simp add: min_of_mono monoI)
        have le_progress_iff[simp]: "(Suc j)  progress σ P' φ (Suc j)  progress σ P' φ (Suc j) = (Suc j)" for φ
          using wf_envs_progress_le[OF MUntil.prems(2), of φ] by auto
        have min_Suc[simp]: "min j (Suc j) = j" by auto
        let ?X = "{i. k. k < Suc j  k min (progress σ P' φ''' (Suc j)) (progress σ P' ψ'' (Suc j))  enat (τ σ k)  enat (τ σ i) + right I}"
        let ?min = "min j (min (progress σ P' φ'' (Suc j)) (progress σ P' ψ'' (Suc j)))"
        have "τ σ ?min  τ σ j"
          by (rule τ_mono) auto
        from m have "?X  {}"
          by (auto dest!: τ_mono[of _ "progress σ P' φ'' (Suc j)" σ]
              simp: not_le not_less φ''' intro!: exI[of _ "progress σ P' φ'' (Suc j)"])
        from m show ?thesis
          using that nts' unfolding wf_ts_def progress.simps
          by (intro cInf_greatest[OF ?X  {}])
            (auto simp: 1 φ''' not_le not_less list_all2_Cons2 upt_eq_Cons_conv less_Suc_eq
              simp del: upt_Suc split: list.splits if_splits
              dest!: spec[of _ "?min"] less_le_trans[of "τ σ i + m" "τ σ _" "τ σ _ + m"] less_τD)
      qed
      moreover have *: "k < progress σ P' ψ (Suc j)" if
        "enat (τ σ i) + right I < enat (case nts' of []  τ σ j | nt # x  nt)"
        "enat (τ σ k - τ σ i)  right I" "ψ = ψ''  ψ = φ''" for k ψ
      proof -
        from that(1,2) obtain m where "right I = enat m"
          "τ σ i + m < (case nts' of []  τ σ j | nt # x  nt)" "τ σ k - τ σ i  m"
          by (cases "right I") auto
        with that(3) nts' progress_le[of σ ψ'' "Suc j"] progress_le[of σ φ'' "Suc j"]
        show ?thesis
          unfolding wf_ts_def le_diff_conv
          by (auto simp: not_le list_all2_Cons2 upt_eq_Cons_conv less_Suc_eq add.commute
              simp del: upt_Suc split: list.splits if_splits dest!: le_less_trans[of "τ σ k"] less_τD)
      qed
      ultimately show ?case using Cons.prems Suc_i_aux'[simplified]
        unfolding a = (t, a1, a2)
        by (auto simp: φ''' 1 sat.simps upt_conv_Cons dest!:  Cons.IH[OF _ aux' Suc_i_aux']
            simp del: upt_Suc split: if_splits prod.splits intro!: iff_exI qtable_cong[OF 3 refl])
    qed
    thm this
    note wf_aux'' = this[OF wf_envs_progress_mono[OF MUntil.prems(2) le_SucI[OF order_refl]]
      wf_auxlist' conjunct2[OF update1, unfolded ne_def length_aux']]
    have "progress σ P (formula.Until φ''' I ψ'') j + length auxlist' =
      progress σ P' (formula.Until φ''' I ψ'') (Suc j) + length auxlist''"
      using wf_aux'' valid_aux'' length_aux'_aux''
      by (auto simp add: ne_def length_aux' length_aux'')
    then have "cur =
      (if progress σ P' (formula.Until φ''' I ψ'') (Suc j) + length auxlist'' = 0 then 0
      else τ σ (progress σ P' (formula.Until φ''' I ψ'') (Suc j) + length auxlist'' - 1))"
      unfolding cur ne_def by auto
    then have "wf_until_aux σ V R args φ'' ψ'' aux'' (progress σ P' (formula.Until φ''' I ψ'') (Suc j)) 
      progress σ P (formula.Until φ''' I ψ'') j + length zs = progress σ P' (formula.Until φ''' I ψ'') (Suc j) 
      progress σ P (formula.Until φ''' I ψ'') j + length zs + length_muaux args aux'' = min (progress σ P' φ''' (Suc j)) (progress σ P' ψ'' (Suc j)) 
      list_all2 (λi. qtable n (fv ψ'') (mem_restr R) (λv. Formula.sat σ V (map the v) i (formula.Until (if pos then φ'' else formula.Neg φ'') I ψ'')))
      [progress σ P (formula.Until φ''' I ψ'') j..<progress σ P (formula.Until φ''' I ψ'') j + length zs] zs"
      using wf_aux'' valid_aux'' fvi_subset
      unfolding wf_until_aux_def length_aux'' args_ivl args_n args_pos by (auto simp only: length_aux'')
  }
  note update = this
  from MUntil.IH(1)[OF φ MUntil.prems(2)] MUntil.IH(2)[OF ψ MUntil.prems(2)] pos pos_eq fvi_subset show ?case
    by (auto 0 4 simp: args_ivl args_n args_pos Until_eq φ''' progress.simps(6) split: prod.split if_splits
        dest!: update[OF refl refl, rotated]
        intro!: wf_mformula.Until[OF _ _ _ _ args_ivl args_n args_L args_R fvi_subset]
        elim!: list.rel_mono_strong qtable_cong
        elim: mbuf2t_take_add'(1)[OF _ wf_envs_P_simps[OF MUntil.prems(2)] buf nts_snoc]
              mbuf2t_take_add'(2)[OF _ wf_envs_P_simps[OF MUntil.prems(2)] buf nts_snoc])
next
  case (MMatchP I mr mrs φs buf nts aux)
  note sat.simps[simp del] mbufnt_take.simps[simp del] mbufn_add.simps[simp del]
  from MMatchP.prems obtain r ψs where eq: "φ' = Formula.MatchP I r"
    and safe: "safe_regex Past Strict r"
    and mr: "to_mregex r = (mr, ψs)"
    and mrs: "mrs = sorted_list_of_set (RPDs mr)"
    and ψs: "list_all2 (wf_mformula σ j P V n R) φs ψs"
    and buf: "wf_mbufn' σ P V j n R r buf"
    and nts: "wf_ts_regex σ P j r nts"
    and aux: "wf_matchP_aux σ V n R I r aux (progress σ P (Formula.MatchP I r) j)"
    by (cases rule: wf_mformula.cases) (auto)
  have nts_snoc: "list_all2 (λi t. t = τ σ i) [progress_regex σ P r j..<Suc j] (nts @ [τ σ j])"
    using nts unfolding wf_ts_regex_def
    by (auto simp add: wf_envs_progress_regex_le[OF MMatchP.prems(2)] intro: list_all2_appendI)
  have update: "wf_matchP_aux σ V n R I r (snd (zs, aux')) (progress σ P' (Formula.MatchP I r) (Suc j)) 
    list_all2 (λi. qtable n (Formula.fv_regex r) (mem_restr R)
      (λv. Formula.sat σ V (map the v) i (Formula.MatchP I r)))
      [progress σ P (Formula.MatchP I r) j..<progress σ P' (Formula.MatchP I r) (Suc j)] (fst (zs, aux'))"
    if eval: "map (fst o meval n (τ σ j) db) φs = xss"
      and eq: "mbufnt_take (λrels t (zs, aux).
        case update_matchP n I mr mrs rels t aux of (z, x)  (zs @ [z], x))
        ([], aux) (mbufn_add xss buf) (nts @ [τ σ j]) = ((zs, aux'), buf', nts')"
    for xss zs aux' buf' nts'
    unfolding progress_simps
  proof (rule mbufnt_take_add_induct'[where j'="Suc j" and z'="(zs, aux')", OF eq wf_envs_P_simps[OF MMatchP.prems(2)] safe mr buf nts_snoc],
      goal_cases xss _ base step)
    case xss
    then show ?case
      using eval ψs
      by (auto simp: list_all3_map map2_map_map list_all3_list_all2_conv list.rel_map
          list.rel_flip[symmetric, of _ ψs φs] dest!: MMatchP.IH(1)[OF _ _ MMatchP.prems(2)]
          elim!: list.rel_mono_strong split: prod.splits)
  next
    case base
    then show ?case
      using aux by auto
  next
    case (step k Xs z)
    then show ?case
      by (auto simp: Un_absorb1 mrs safe mr elim!: update_matchP(1) list_all2_appendI
          dest!: update_matchP(2) split: prod.split)
  qed simp
  then show ?case using ψs
    by (auto simp: eq mr mrs safe map_split_alt list.rel_flip[symmetric, of _ ψs φs]
        list_all3_map map2_map_map list_all3_list_all2_conv list.rel_map intro!: wf_mformula.intros
        elim!: list.rel_mono_strong mbufnt_take_add'(1)[OF _ wf_envs_P_simps[OF MMatchP.prems(2)] safe mr buf nts_snoc]
        mbufnt_take_add'(2)[OF _ wf_envs_P_simps[OF MMatchP.prems(2)] safe mr buf nts_snoc]
        dest!: MMatchP.IH[OF _ _ MMatchP.prems(2)] split: prod.splits)
next
  case (MMatchF I mr mrs φs buf nts aux)
  note sat.simps[simp del] mbufnt_take.simps[simp del] mbufn_add.simps[simp del] progress_simps[simp del]
  from MMatchF.prems obtain r ψs where eq: "φ' = Formula.MatchF I r"
    and safe: "safe_regex Futu Strict r"
    and mr: "to_mregex r = (mr, ψs)"
    and mrs: "mrs = sorted_list_of_set (LPDs mr)"
    and ψs: "list_all2 (wf_mformula σ j P V n R) φs ψs"
    and buf: "wf_mbufn' σ P V j n R r buf"
    and nts: "wf_ts_regex σ P j r nts"
    and aux: "wf_matchF_aux σ V n R I r aux (progress σ P (Formula.MatchF I r) j) 0"
    and length_aux: "progress σ P (Formula.MatchF I r) j + length aux = progress_regex σ P r j"
    by (cases rule: wf_mformula.cases) (auto)
  have nts_snoc: "list_all2 (λi t. t = τ σ i)
    [progress_regex σ P r j..<Suc j] (nts @ [τ σ j])"
    using nts unfolding wf_ts_regex_def
    by (auto simp add: wf_envs_progress_regex_le[OF MMatchF.prems(2)] intro: list_all2_appendI)
  {
    fix xss zs aux' aux'' buf' nts'
    assume eval: "map (fst o meval n (τ σ j) db) φs = xss"
      and eq1: "mbufnt_take (update_matchF n I mr mrs) aux (mbufn_add xss buf) (nts @ [τ σ j]) =
        (aux', buf', nts')"
      and eq2: "eval_matchF I mr (case nts' of []  τ σ j | nt # _  nt) aux' = (zs, aux'')"
    have update1: "wf_matchF_aux σ V n R I r aux' (progress σ P (Formula.MatchF I r) j) 0 
      progress σ P (Formula.MatchF I r) j + length aux' = progress_regex σ P' r (Suc j)"
      using eval nts_snoc nts_snoc length_aux aux ψs
      by (elim mbufnt_take_add_induct'[where j'="Suc j", OF eq1 wf_envs_P_simps[OF MMatchF.prems(2)] safe mr buf])
        (auto simp: length_update_matchF
          list_all3_map map2_map_map list_all3_list_all2_conv list.rel_map list.rel_flip[symmetric, of _ ψs φs]
          dest!: MMatchF.IH[OF _ _ MMatchF.prems(2)]
          elim: wf_update_matchF[OF _ safe mr mrs] elim!: list.rel_mono_strong)
    from MMatchF.prems(2) have nts': "wf_ts_regex σ P' (Suc j) r nts'"
      using eval eval nts_snoc ψs
      unfolding wf_ts_regex_def
      by (intro mbufnt_take_eqD(2)[OF eq1 wf_mbufn_add[where js'="map (λφ. progress σ P' φ (Suc j)) ψs",
              OF buf[unfolded wf_mbufn'_def mr prod.case]]])
        (auto simp: to_mregex_progress[OF safe mr] Mini_def
          list_all3_map map2_map_map list_all3_list_all2_conv list.rel_map list.rel_flip[symmetric, of _ ψs φs]
          list_all2_Cons1 elim!: list.rel_mono_strong intro!: list.rel_refl_strong
          dest!: MMatchF.IH[OF _ _ MMatchF.prems(2)])
    have "i  progress σ P' (Formula.MatchF I r) (Suc j) 
      wf_matchF_aux σ V n R I r aux' i 0 
      i + length aux' = progress_regex σ P' r (Suc j) 
      wf_matchF_aux σ V n R I r aux'' (progress σ P' (Formula.MatchF I r) (Suc j)) 0 
        i + length zs = progress σ P' (Formula.MatchF I r) (Suc j) 
        i + length zs + length aux'' = progress_regex σ P' r (Suc j) 
        list_all2 (λi. qtable n (Formula.fv_regex r) (mem_restr R)
          (λv. Formula.sat σ V (map the v) i (Formula.MatchF I r)))
          [i..<i + length zs] zs" for i
      using eq2
    proof (induction aux' arbitrary: zs aux'' i)
      case Nil
      then show ?case by (auto dest!: antisym[OF progress_MatchF_le])
    next
      case (Cons a aux')
      obtain t rels rel where "a = (t, rels, rel)" by (cases a)
      from Cons.prems(2) have aux': "wf_matchF_aux σ V n R I r aux' (Suc i) 0"
        by (rule wf_matchF_aux_Cons)
      from Cons.prems(2) have 1: "t = τ σ i"
        unfolding a = (t, rels, rel) by (rule wf_matchF_aux_Cons1)
      from Cons.prems(2) have 3: "qtable n (Formula.fv_regex r) (mem_restr R) (λv.
        (ji. j < Suc (i + length aux')  mem (τ σ j - τ σ i) I  Regex.match (Formula.sat σ V (map the v)) r i j)) rel"
        unfolding a = (t, rels, rel) using wf_matchF_aux_Cons3 by force
      from Cons.prems(3) have Suc_i_aux': "Suc i + length aux' = progress_regex σ P' r (Suc j)"
        by simp
      have "i  progress σ P' (Formula.MatchF I r) (Suc j)"
        if "enat (case nts' of []  τ σ j | nt # x  nt)  enat t + right I"
        using that nts' unfolding wf_ts_regex_def progress_simps
        by (auto simp add: 1 list_all2_Cons2 upt_eq_Cons_conv
            intro!: cInf_lower τ_mono elim!: order.trans[rotated] simp del: upt_Suc split: if_splits list.splits)
      moreover
      have "Suc i  progress σ P' (Formula.MatchF I r) (Suc j)"
        if "enat t + right I < enat (case nts' of []  τ σ j | nt # x  nt)"
      proof -
        from that obtain m where m: "right I = enat m" by (cases "right I") auto
        have τ_min:  "τ σ (min j k) = min (τ σ j) (τ σ k)" for k
          by (simp add: min_of_mono monoI)
        have le_progress_iff[simp]: "Suc j  progress σ P' φ (Suc j)  progress σ P' φ (Suc j) = (Suc j)" for φ
          using wf_envs_progress_le[OF MMatchF.prems(2), of φ] by auto
        have min_Suc[simp]: "min j (Suc j) = j" by auto
        let ?X = "{i. k. k < Suc j  k  progress_regex σ P' r (Suc j)  enat (τ σ k)  enat (τ σ i) + right I}"
        let ?min = "min j (progress_regex σ P' r (Suc j))"
        have "τ σ ?min  τ σ j"
          by (rule τ_mono) auto
        from m have "?X  {}"
          by (auto dest!: less_τD add_lessD1 simp: not_le not_less)
        from m show ?thesis
          using that nts' wf_envs_progress_regex_le[OF MMatchF.prems(2), of r]
          unfolding wf_ts_regex_def progress_simps
          by (intro cInf_greatest[OF ?X  {}])
            (auto simp: 1 not_le not_less list_all2_Cons2 upt_eq_Cons_conv less_Suc_eq
              simp del: upt_Suc split: list.splits if_splits
              dest!: spec[of _ "?min"] less_le_trans[of "τ σ i + m" "τ σ _" "τ σ _ + m"] less_τD)
      qed
      moreover have *: "k < progress_regex σ P' r (Suc j)" if
        "enat (τ σ i) + right I < enat (case nts' of []  τ σ j | nt # x  nt)"
        "enat (τ σ k - τ σ i)  right I" for k
      proof -
        from that(1,2) obtain m where "right I = enat m"
          "τ σ i + m < (case nts' of []  τ σ j | nt # x  nt)" "τ σ k - τ σ i  m"
          by (cases "right I") auto
        with nts' wf_envs_progress_regex_le[OF MMatchF.prems(2), of r]
        show ?thesis
          unfolding wf_ts_regex_def le_diff_conv
          by (auto simp: not_le list_all2_Cons2 upt_eq_Cons_conv less_Suc_eq add.commute
              simp del: upt_Suc split: list.splits if_splits dest!: le_less_trans[of "τ σ k"] less_τD)
      qed
      ultimately show ?case using Cons.prems Suc_i_aux'[simplified]
        unfolding a = (t, rels, rel)
        by (auto simp: 1 sat.simps upt_conv_Cons dest!:  Cons.IH[OF _ aux' Suc_i_aux']
            simp del: upt_Suc split: if_splits prod.splits intro!: iff_exI qtable_cong[OF 3 refl])

    qed
    note this[OF progress_mono_gen[OF le_SucI, OF order.refl] conjunct1[OF update1] conjunct2[OF update1]]
  }
  note update = this[OF refl, rotated]
  with MMatchF.prems(2) show ?case using ψs
    by (auto simp: eq mr mrs safe map_split_alt list.rel_flip[symmetric, of _ ψs φs]
        list_all3_map map2_map_map list_all3_list_all2_conv list.rel_map intro!: wf_mformula.intros
        elim!: list.rel_mono_strong mbufnt_take_add'(1)[OF _ wf_envs_P_simps[OF MMatchF.prems(2)] safe mr buf nts_snoc]
        mbufnt_take_add'(2)[OF _ wf_envs_P_simps[OF MMatchF.prems(2)] safe mr buf nts_snoc]
        dest!: MMatchF.IH[OF _ _ MMatchF.prems(2)] update split: prod.splits)
qed


subsubsection ‹Monitor step›

lemma (in maux) wf_mstate_mstep: "wf_mstate φ π R st  last_ts π  snd tdb 
  wf_mstate φ (psnoc π tdb) R (snd (mstep (map_prod mk_db id tdb) st))"
  unfolding wf_mstate_def mstep_def Let_def
  by (fastforce simp add: progress_mono le_imp_diff_is_add split: prod.splits
      elim!: prefix_of_psnocE dest: meval[OF _ wf_envs_mk_db] list_all2_lengthD)

definition "flatten_verdicts Vs = ( (set (map (λ(i, X). (λv. (i, v)) ` X) Vs)))"

lemma flatten_verdicts_append[simp]:
  "flatten_verdicts (Vs @ Us) = flatten_verdicts Vs  flatten_verdicts Us"
  by (induct Vs) (auto simp: flatten_verdicts_def)

lemma (in maux) mstep_output_iff:
  assumes "wf_mstate φ π R st" "last_ts π  snd tdb" "prefix_of (psnoc π tdb) σ" "mem_restr R v"
  shows "(i, v)  flatten_verdicts (fst (mstep (map_prod mk_db id tdb) st)) 
    progress σ Map.empty φ (plen π)  i  i < progress σ Map.empty φ (Suc (plen π)) 
    wf_tuple (Formula.nfv φ) (Formula.fv φ) v  Formula.sat σ Map.empty (map the v) i φ"
proof -
  from prefix_of_psnocE[OF assms(3,2)] have "prefix_of π σ"
    "Γ σ (plen π) = fst tdb" "τ σ (plen π) = snd tdb" by auto
  moreover from assms(1) prefix_of π σ have "mstate_n st = Formula.nfv φ"
    "mstate_i st = progress σ Map.empty φ (plen π)" "wf_mformula σ (plen π) Map.empty Map.empty (mstate_n st) R (mstate_m st) φ"
    unfolding wf_mstate_def by blast+
  moreover from meval[OF wf_mformula σ (plen π) Map.empty Map.empty (mstate_n st) R (mstate_m st) φ wf_envs_mk_db] obtain Vs st' where
    "meval (mstate_n st) (τ σ (plen π)) (mk_db (Γ σ (plen π))) (mstate_m st) = (Vs, st')"
    "wf_mformula σ (Suc (plen π))  Map.empty Map.empty (mstate_n st) R st' φ"
    "list_all2 (λi. qtable (mstate_n st) (fv φ) (mem_restr R) (λv. Formula.sat σ Map.empty (map the v) i φ))
      [progress σ Map.empty φ (plen π)..<progress σ Map.empty φ (Suc (plen π))] Vs" by blast
  moreover from this assms(4) have "qtable (mstate_n st) (fv φ) (mem_restr R)
    (λv. Formula.sat σ Map.empty (map the v) i φ) (Vs ! (i - progress σ Map.empty φ (plen π)))"
    if "progress σ Map.empty φ (plen π)  i" "i < progress σ Map.empty φ (Suc (plen π))"
    using that by (auto simp: list_all2_conv_all_nth
        dest!: spec[of _ "(i - progress σ Map.empty φ (plen π))"])
  ultimately show ?thesis
    using assms(4) unfolding mstep_def Let_def flatten_verdicts_def
    by (auto simp: in_set_enumerate_eq list_all2_conv_all_nth progress_mono le_imp_diff_is_add
        elim!: in_qtableE in_qtableI intro!: bexI[of _ "(i, Vs ! (i - progress σ Map.empty φ (plen π)))"])
qed


subsubsection ‹Monitor function›

locale verimon = verimon_spec + maux

lemma (in verimon) mstep_mverdicts:
  assumes wf: "wf_mstate φ π R st"
    and le[simp]: "last_ts π  snd tdb"
    and restrict: "mem_restr R v"
  shows "(i, v)  flatten_verdicts (fst (mstep (map_prod mk_db id tdb) st)) 
    (i, v)  M (psnoc π tdb) - M π"
proof -
  obtain σ where p2: "prefix_of (psnoc π tdb) σ"
    using ex_prefix_of by blast
  with le have p1: "prefix_of π σ" by (blast elim!: prefix_of_psnocE)
  show ?thesis
    unfolding M_def
    by (auto 0 3 simp: p2 progress_prefix_conv[OF _ p1] sat_prefix_conv[OF _ p1] not_less
        pprogress_eq[OF p1] pprogress_eq[OF p2]
        dest: mstep_output_iff[OF wf le p2 restrict, THEN iffD1] spec[of _ σ]
        mstep_output_iff[OF wf le _ restrict, THEN iffD1] progress_sat_cong[OF p1]
        intro: mstep_output_iff[OF wf le p2 restrict, THEN iffD2] p1)
qed

context maux
begin

primrec msteps0 where
  "msteps0 [] st = ([], st)"
| "msteps0 (tdb # π) st =
    (let (V', st') = mstep (map_prod mk_db id tdb) st; (V'', st'') = msteps0 π st' in (V' @ V'', st''))"

primrec msteps0_stateless where
  "msteps0_stateless [] st = []"
| "msteps0_stateless (tdb # π) st = (let (V', st') = mstep (map_prod mk_db id tdb) st in V' @ msteps0_stateless π st')"

lemma msteps0_msteps0_stateless: "fst (msteps0 w st) = msteps0_stateless w st"
  by (induct w arbitrary: st) (auto simp: split_beta)

lift_definition msteps :: "Formula.prefix  ('msaux, 'muaux) mstate  (nat × event_data table) list × ('msaux, 'muaux) mstate"
  is msteps0 .

lift_definition msteps_stateless :: "Formula.prefix  ('msaux, 'muaux) mstate  (nat × event_data table) list"
  is msteps0_stateless .

lemma msteps_msteps_stateless: "fst (msteps w st) = msteps_stateless w st"
  by transfer (rule msteps0_msteps0_stateless)

lemma msteps0_snoc: "msteps0 (π @ [tdb]) st =
   (let (V', st') = msteps0 π st; (V'', st'') = mstep (map_prod mk_db id tdb) st' in (V' @ V'', st''))"
  by (induct π arbitrary: st) (auto split: prod.splits)

lemma msteps_psnoc: "last_ts π  snd tdb  msteps (psnoc π tdb) st =
   (let (V', st') = msteps π st; (V'', st'') = mstep (map_prod mk_db id tdb) st' in (V' @ V'', st''))"
  by transfer' (auto simp: msteps0_snoc split: list.splits prod.splits if_splits)

definition monitor where
  "monitor φ π = msteps_stateless π (minit_safe φ)"

end

lemma Suc_length_conv_snoc: "(Suc n = length xs) = (y ys. xs = ys @ [y]  length ys = n)"
  by (cases xs rule: rev_cases) auto

lemma (in verimon) wf_mstate_msteps: "wf_mstate φ π R st  mem_restr R v  π  π' 
  X = msteps (pdrop (plen π) π') st  wf_mstate φ π' R (snd X) 
  ((i, v)  flatten_verdicts (fst X)) = ((i, v)  M π' - M π)"
proof (induct "plen π' - plen π" arbitrary: X st π π')
  case 0
  from 0(1,4,5) have "π = π'"  "X = ([], st)"
    by (transfer; auto)+
  with 0(2) show ?case unfolding flatten_verdicts_def by simp
next
  case (Suc x)
  from Suc(2,5) obtain π'' tdb where "x = plen π'' - plen π"  "π  π''"
    "π' = psnoc π'' tdb" "pdrop (plen π) (psnoc π'' tdb) = psnoc (pdrop (plen π) π'') tdb"
    "last_ts (pdrop (plen π) π'')  snd tdb" "last_ts π''  snd tdb"
    "π''  psnoc π'' tdb"
  proof (atomize_elim, transfer, elim exE, goal_cases prefix)
    case (prefix _ _ π' _ π_tdb)
    then show ?case
    proof (cases π_tdb rule: rev_cases)
      case (snoc π tdb)
      with prefix show ?thesis
        by (intro bexI[of _ "π' @ π"] exI[of _ tdb])
          (force simp: sorted_append append_eq_Cons_conv split: list.splits if_splits)+
    qed simp
  qed
  with Suc(1)[OF this(1) Suc.prems(1,2) this(2) refl] Suc.prems show ?case
    unfolding msteps_msteps_stateless[symmetric]
    by (auto simp: msteps_psnoc split_beta mstep_mverdicts
        dest: mono_monitor[THEN set_mp, rotated] intro!: wf_mstate_mstep)
qed

lemma (in verimon) wf_mstate_msteps_stateless:
  assumes "wf_mstate φ π R st" "mem_restr R v" "π  π'"
  shows "(i, v)  flatten_verdicts (msteps_stateless (pdrop (plen π) π') st)  (i, v)  M π' - M π"
  using wf_mstate_msteps[OF assms refl] unfolding msteps_msteps_stateless by simp

lemma (in verimon) wf_mstate_msteps_stateless_UNIV: "wf_mstate φ π UNIV st  π  π' 
  flatten_verdicts (msteps_stateless (pdrop (plen π) π') st) = M π' - M π"
  by (auto dest: wf_mstate_msteps_stateless[OF _ mem_restr_UNIV])

lemma (in verimon) mverdicts_Nil: "M pnil = {}"
  by (simp add: M_def pprogress_eq)

context maux
begin

lemma minit_safe_minit: "mmonitorable φ  minit_safe φ = minit φ"
  unfolding minit_safe_def monitorable_formula_code by simp

lemma wf_mstate_minit_safe: "mmonitorable φ  wf_mstate φ pnil R (minit_safe φ)"
  using wf_mstate_minit minit_safe_minit mmonitorable_def by metis

end

lemma (in verimon) monitor_mverdicts: "flatten_verdicts (monitor φ π) = M π"
  unfolding monitor_def using monitorable
  by (subst wf_mstate_msteps_stateless_UNIV[OF wf_mstate_minit_safe, simplified])
    (auto simp: mmonitorable_def mverdicts_Nil)

subsection ‹Collected correctness results›

context verimon
begin

text ‹We summarize the main results proved above.
\begin{enumerate}
\item The term @{term M} describes semantically the monitor's expected behaviour:
\begin{itemize}
\item @{thm[source] mono_monitor}: @{thm mono_monitor[no_vars]}
\item @{thm[source] sound_monitor}: @{thm sound_monitor[no_vars]}
\item @{thm[source] complete_monitor}: @{thm complete_monitor[no_vars]}
\item @{thm[source] sliceable_M}: @{thm sliceable_M[no_vars]}
\end{itemize}
\item The executable monitor's online interface @{term minit_safe} and @{term mstep}
  preserves the invariant @{term wf_mstate} and produces the the verdicts according
  to @{term M}:
\begin{itemize}
\item @{thm[source] wf_mstate_minit_safe}: @{thm wf_mstate_minit_safe[no_vars]}
\item @{thm[source] wf_mstate_mstep}: @{thm wf_mstate_mstep[no_vars]}
\item @{thm[source] mstep_mverdicts}: @{thm mstep_mverdicts[no_vars]}
\end{itemize}
\item The executable monitor's offline interface @{term monitor} implements @{term M}:
\begin{itemize}
\item @{thm[source] monitor_mverdicts}: @{thm monitor_mverdicts[no_vars]}
\end{itemize}
\end{enumerate}
›

end

(*<*)
end
(*>*)