Theory Tree_Automata_Abstract_Impl
theory Tree_Automata_Abstract_Impl
  imports Tree_Automata_Det Horn_Fset
begin
section βΉComputing state derivationβΊ
lemma ta_der_Var_code [code]:
  "ta_der π (Var q) = finsert q ((eps π)|β§+| |``| {|q|})"
  by (auto)
lemma ta_der_Fun_code [code]:
  "ta_der π (Fun f ts) =
     (let args = map (ta_der π) ts in
      let P = (Ξ» r. case r of TA_rule g ps p β f = g β§ list_all2 fmember ps args) in
      let S = r_rhs |`| ffilter P (rules π) in
         S |βͺ| (eps π)|β§+| |``| S)" (is "?Ls = ?Rs")
proof
  {fix q assume "q |β| ?Ls" then have "q |β| ?Rs"
      apply (simp add: Let_def fImage_iff fBex_def image_iff)
      by (smt (verit, ccfv_threshold) IntI length_map list_all2_conv_all_nth mem_Collect_eq nth_map
          ta_rule.case ta_rule.sel(3))
    }
  then show "?Ls |β| ?Rs" by blast
next
  {fix q assume "q |β| ?Rs" then have "q |β| ?Ls"
      apply (auto simp: Let_def ffmember_filter fimage_iff fBex_def list_all2_conv_all_nth fImage_iff
                  split!: ta_rule.splits)
      apply (metis ta_rule.collapse)
      apply blast
      done}
  then show "?Rs |β| ?Ls" by blast
qed
definition eps_free_automata where
  "eps_free_automata epscl π =
  (let ruleps = (Ξ» r. finsert (r_rhs r) (epscl |``| {|r_rhs r|})) in
   let rules = (Ξ» r. (Ξ» q. TA_rule (r_root r) (r_lhs_states r) q) |`| (ruleps r)) |`| (rules π) in
   TA ( |β| rules) {||})"
lemma eps_free [code]:
  "eps_free π = eps_free_automata ((eps π)|β§+|) π"
  apply (intro TA_equalityI)
   apply (auto simp: eps_free_def eps_free_rulep_def eps_free_automata_def)
  using fBex_def apply fastforce
  apply (metis ta_rule.exhaust_sel)+
  done
lemma eps_of_eps_free_automata [simp]:
  "eps (eps_free_automata S π) = {||}"
  by (auto simp add: eps_free_automata_def)
lemma eps_free_automata_empty [simp]:
  "eps π = {||} βΉ eps_free_automata {||} π = π"
  by (auto simp add: eps_free_automata_def intro!: TA_equalityI)
section βΉComputing the restriction of tree automata to state setβΊ
lemma ta_restrict [code]:
  "ta_restrict π Q =
     (let rules =  ffilter (Ξ» r. case r of TA_rule f ps p β fset_of_list ps |β| Q β§ p |β| Q) (rules π) in
      let eps = ffilter (Ξ» r. case r of (p, q) β p |β| Q β§ q |β| Q) (eps π) in
      TA rules eps)"
  by (auto simp: Let_def ta_restrict_def split!: ta_rule.splits intro: finite_subset[OF _ finite_Collect_ta_rule])
section βΉComputing the epsilon transition for the product automatonβΊ
lemma prod_eps[code_unfold]:
  "fCollect (prod_epsLp π β¬) = (Ξ» ((p, q), r). ((p, r), (q, r))) |`| (eps π |Γ| π¬ β¬)"
  "fCollect (prod_epsRp π β¬) = (Ξ» ((p, q), r). ((r, p), (r, q))) |`| (eps β¬ |Γ| π¬ π)"
  by (auto simp: finite_prod_epsLp prod_epsLp_def finite_prod_epsRp prod_epsRp_def image_iff
      fSigma.rep_eq) force+
section βΉComputing reachabilityβΊ
inductive_set ta_reach for π where
   rule [intro]: "f qs β q |β| rules π βΉ β i < length qs. qs ! i β ta_reach π βΉ q β ta_reach π"
 |  eps [intro]: "q β ta_reach π βΉ (q, r) |β| eps π βΉ r β ta_reach π"
lemma ta_reach_eps_transI:
  assumes "(p, q) |β| (eps π)|β§+|" "p β ta_reach π"
  shows "q β ta_reach π" using assms
  by (induct rule: ftrancl_induct) auto
lemma ta_reach_ground_term_der:
  assumes "q β ta_reach π"
  shows "β t. ground t β§ q |β| ta_der π t" using assms
proof (induct)
  case (rule f qs q)
  then obtain ts where "length ts = length qs"
    "β i < length qs. ground (ts ! i)"
    "β i < length qs. qs ! i |β| ta_der π (ts ! i)"
    using Ex_list_of_length_P[of "length qs" "Ξ» t i. ground t β§ qs ! i |β| ta_der π t"]
    by auto
  then show ?case using rule(1)
    by (auto dest!: in_set_idx intro!: exI[of _ "Fun f ts"]) blast
qed (auto, meson ta_der_eps)
lemma ground_term_der_ta_reach:
  assumes "ground t" "q |β| ta_der π t"
  shows "q β ta_reach π" using assms(2, 1)
  by (induct rule: ta_der_induct) (auto simp add: rule ta_reach_eps_transI)
lemma ta_reach_reachable:
  "ta_reach π = fset (ta_reachable π)"
  using ta_reach_ground_term_der[of _ π]
  using ground_term_der_ta_reach[of _ _ π]
  unfolding ta_reachable_def
  by auto
subsection βΉHorn setup for reachable statesβΊ
definition "reach_rules π =
  {qs ββ©h q | f qs q. TA_rule f qs q |β| rules π} βͺ
  {[q] ββ©h r | q r. (q, r) |β| eps π}"
locale reach_horn =
  fixes π :: "('q, 'f) ta"
begin
sublocale horn "reach_rules π" .
lemma reach_infer0: "infer0 = {q | f q. TA_rule f [] q |β| rules π}"
  by (auto simp: horn.infer0_def reach_rules_def)
lemma reach_infer1:
  "infer1 p X = {r | f qs r. TA_rule f qs r |β| rules π β§ p β set qs β§ set qs β insert p X} βͺ
   {r | r. (p, r) |β| eps π}"
  unfolding reach_rules_def
  by (auto simp: horn.infer1_def simp flip: ex_simps(1))
lemma reach_sound:
  "ta_reach π = saturate"
proof (intro set_eqI iffI, goal_cases lr rl)
  case (lr x) obtain p where x: "p = ta_reach π" by auto
  show ?case using lr unfolding x
  proof (induct)
    case (rule f qs q)
    then show ?case
      by (intro infer[of qs q]) (auto simp: reach_rules_def dest: in_set_idx)
  next
    case (eps q r)
    then show ?case
      by (intro infer[of "[q]" r]) (auto simp: reach_rules_def)
  qed
next
  case (rl x)
  then show ?case
    by (induct) (auto simp: reach_rules_def)
qed
end
subsection βΉComputing productivityβΊ
text βΉFirst, use an alternative definition of productivityβΊ
inductive_set ta_productive_ind :: "'q fset β ('q,'f) ta β 'q set" for P and π :: "('q,'f) ta" where
  basic [intro]: "q |β| P βΉ q β ta_productive_ind P π"
| eps [intro]: "(p, q) |β| (eps π)|β§+| βΉ q β ta_productive_ind P π βΉ p β ta_productive_ind P π"
| rule: "TA_rule f qs q |β| rules π βΉ q β ta_productive_ind P π βΉ q' β set qs βΉ q' β ta_productive_ind P π"
lemma ta_productive_ind:
  "ta_productive_ind P π = fset (ta_productive P π)" (is "?LS = ?RS")
proof -
  {fix q assume "q β ?LS" then have "q β ?RS"
      by (induct) (auto dest: ta_prod_epsD intro: ta_productive_setI,
         metis (full_types) in_set_conv_nth rule_reachable_ctxt_exist ta_productiveI')}
  moreover
  {fix q assume "q β ?RS" note * = this
    from ta_productiveE[OF *] obtain r C where
      reach : "r |β| ta_der π Cβ¨Var qβ©" and f: "r |β| P" by auto
    from f have "r β ta_productive_ind P π" "r |β| ta_productive P π"
      by (auto intro: ta_productive_setI)
    then have "q β ?LS" using reach
    proof (induct C arbitrary: q r)
      case (More f ss C ts)
      from iffD1 ta_der_Fun[THEN iffD1, OF More(4)[unfolded intp_actxt.simps]] obtain ps p where
        inv: "f ps β p |β| rules π" "p = r β¨ (p, r) |β| (eps π)|β§+|" "length ps = length (ss @ Cβ¨Var qβ© # ts)"
             "ps ! length ss |β| ta_der π Cβ¨Var qβ©"
        by (auto simp: nth_append_Cons split: if_splits)
      then have "p β ta_productive_ind P π βΉ p |β| ta_der π Cβ¨Var qβ© βΉ q β ta_productive_ind P π" for p
        using More(1) calculation by auto
      note [intro!] = this[of "ps ! length ss"]
      show ?case using More(2) inv
        by (auto simp: nth_append_Cons ta_productive_ind.rule)
           (metis less_add_Suc1 nth_mem ta_productive_ind.simps)
    qed (auto intro: ta_productive_setI)
  }
  ultimately show ?thesis by auto
qed
subsubsection βΉHorn setup for productive statesβΊ
definition "productive_rules P π = {[] ββ©h q | q. q |β| P} βͺ
  {[r] ββ©h q | q r. (q, r) |β| eps π} βͺ
  {[q] ββ©h r | f qs q r. TA_rule f qs q |β| rules π β§ r β set qs}"
locale productive_horn =
  fixes π :: "('q, 'f) ta" and P :: "'q fset"
begin
sublocale horn "productive_rules P π" .
lemma productive_infer0: "infer0 = fset P"
  by (auto simp: productive_rules_def horn.infer0_def)
lemma productive_infer1:
  "infer1 p X = {r | r. (r, p) |β| eps π} βͺ
    {r | f qs r. TA_rule f qs p |β| rules π β§ r β set qs}"
  unfolding productive_rules_def horn_infer1_union
  by (auto simp add: horn.infer1_def)
     (metis insertCI list.set(1) list.simps(15) singletonD subsetI)
lemma productive_sound:
  "ta_productive_ind P π = saturate"
proof (intro set_eqI iffI, goal_cases lr rl)
  case (lr p) then show ?case using lr
  proof (induct)
    case (basic q)
    then show ?case
      by (intro infer[of "[]" q]) (auto simp: productive_rules_def)
  next
    case (eps p q) then show ?case
    proof (induct rule: ftrancl_induct)
      case (Base p q)
      then show ?case using infer[of "[q]" p]
        by (auto simp: productive_rules_def)
    next
      case (Step p q r)
      then show ?case using infer[of "[r]" q]
        by (auto simp: productive_rules_def)
    qed
  next
    case (rule f qs q p)
    then show ?case
      by (intro infer[of "[q]" p]) (auto simp: productive_rules_def)
  qed
next
  case (rl p)
  then show ?case
    by (induct) (auto simp: productive_rules_def ta_productive_ind.rule)
qed
end
subsection βΉHorn setup for power set construction statesβΊ
lemma prod_list_exists:
  assumes "fst p β set qs" "set qs β insert (fst p) (fst ` X)"
  obtains as where "p β set as" "map fst as = qs" "set as β insert p X"
proof -
  from assms have "qs β lists (fst ` (insert p X))" by blast
  then obtain ts where ts: "map fst ts = qs" "ts β lists (insert p X)"
    by (metis image_iff lists_image)
  then obtain i where mem: "i < length qs" "qs ! i = fst p" using assms(1)
    by (metis in_set_idx)
  from ts have p: "ts[i := p] β lists (insert p X)"
    using set_update_subset_insert by fastforce
  then have "p β set (ts[i := p])" "map fst (ts[i := p]) = qs" "set (ts[i := p]) β insert p X"
    using mem ts(1)
    by (auto simp add: nth_list_update set_update_memI intro!: nth_equalityI)
  then show ?thesis using that
    by blast
qed
definition "ps_states_rules π = {rs ββ©h (Wrapp q) | rs f q.
    q = ps_reachable_states π f (map ex rs) β§ q β  {||}}"
locale ps_states_horn =
  fixes π :: "('q, 'f) ta"
begin
sublocale horn "ps_states_rules π" .
lemma ps_construction_infer0: "infer0 =
  {Wrapp q | f q. q = ps_reachable_states π f [] β§ q β  {||}}"
    by (auto simp: ps_states_rules_def horn.infer0_def)
lemma ps_construction_infer1:
  "infer1 p X = {Wrapp q | f qs q. q = ps_reachable_states π f (map ex qs) β§ q β  {||} β§
   p β set qs β§ set qs β insert p X}"
  unfolding ps_states_rules_def horn_infer1_union
  by (auto simp add: horn.infer1_def ps_reachable_states_def comp_def elim!: prod_list_exists) 
lemma ps_states_sound:
  "ps_states_set π = saturate"
proof (intro set_eqI iffI, goal_cases lr rl)
  case (lr p) then show ?case using lr
  proof (induct)
    case (1 ps f)
    then have "ps ββ©h (Wrapp (ps_reachable_states π f (map ex ps))) β ps_states_rules π"
      by (auto simp: ps_states_rules_def)
    then show ?case using horn.saturate.simps 1
      by fastforce
  qed
next
  case (rl p)
  then obtain q where "q β saturate" "q = p" by blast
  then show ?case
    by (induct arbitrary: p)
       (auto simp: ps_states_rules_def intro!: ps_states_set.intros)
qed
end
definition ps_reachable_states_cont where
  "ps_reachable_states_cont Ξ Ξβ©Ξ΅ f ps =
   (let R = ffilter (Ξ» r. case r of TA_rule g qs q β f = g β§ list_all2 (|β|) qs ps) Ξ in
    let S = r_rhs |`| R in
    S |βͺ| Ξβ©Ξ΅|β§+| |``| S)"
lemma ps_reachable_states [code]:
  "ps_reachable_states (TA Ξ Ξβ©Ξ΅) f ps = ps_reachable_states_cont Ξ Ξβ©Ξ΅ f ps"
  by (auto simp: ps_reachable_states_fmember ps_reachable_states_cont_def Let_def fimage_iff fBex_def
           split!: ta_rule.splits) force+
definition ps_rules_cont where
 "ps_rules_cont π Q =
   (let sig = ta_sig π in
    let qss = (Ξ» (f, n). (f, n, fset_of_list (List.n_lists n (sorted_list_of_fset Q)))) |`| sig in
    let res = (Ξ» (f, n, Qs). (Ξ» qs. TA_rule f qs (Wrapp (ps_reachable_states π f (map ex qs)))) |`| Qs) |`| qss in
      ffilter (Ξ» r. ex (r_rhs r) β  {||}) ( |β| res))"
lemma ps_rules [code]:
  "ps_rules π Q = ps_rules_cont π Q"
  using ps_reachable_states_sig finite_ps_rulesp_unfolded[of Q π]
  unfolding ps_rules_cont_def
  apply (auto simp: fset_of_list_elem ps_rules_def fin_mono ps_rulesp_def
    image_iff set_n_lists split!: prod.splits dest!: in_set_idx)
  by fastforce+
end