Theory Lazy_Intruder

(*  Title:      Lazy_Intruder.thy
    Author:     Andreas Viktor Hess, DTU
    SPDX-License-Identifier: BSD-3-Clause
*)

section ‹The Lazy Intruder›
theory Lazy_Intruder
imports Strands_and_Constraints Intruder_Deduction
begin

context intruder_model
begin

subsection ‹Definition of the Lazy Intruder›
text ‹The lazy intruder constraint reduction system, defined as a relation on constraint states›
inductive_set LI_rel::
    "((('fun,'var) strand × (('fun,'var) subst)) ×
       ('fun,'var) strand × (('fun,'var) subst)) set"
  and LI_rel' (infix "" 50)
  and LI_rel_trancl (infix "+" 50)
  and LI_rel_rtrancl (infix "*" 50)
where
  "A  B  (A,B)  LI_rel"
| "A + B  (A,B)  LI_rel+"
| "A * B  (A,B)  LI_rel*"

| Compose: "simple S; length T = arity f; public f
             (S@Send [Fun f T]#S',θ)  (S@(map Send1 T)@S',θ)"
| Unify: "simple S; Fun f T'  ikst S; Some δ = mgu (Fun f T) (Fun f T')
           (S@Send [Fun f T]#S',θ)  ((S@S') st δ,θ s δ)"
| Equality: "simple S; Some δ = mgu t t'
           (S@Equality _ t t'#S',θ)  ((S@S') st δ,θ s δ)"


text ‹A "pre-processing step" to be applied before constraint reduction. It transforms constraints
such that exactly one message is transmitted in each message transmission step. It is sound and
complete and preserves the various well-formedness properties required by the lazy intruder.›
fun LI_preproc where
  "LI_preproc [] = []"
| "LI_preproc (Send ts#S) = map Send1 ts@LI_preproc S"
| "LI_preproc (Receive ts#S) = map Receive1 ts@LI_preproc S"
| "LI_preproc (x#S) = x#LI_preproc S"

definition LI_preproc_prop where
  "LI_preproc_prop S  ts. Send ts  set S  Receive ts  set S  (t. ts = [t])" 


subsection ‹Lemmata: Preprocessing ›
lemma LI_preproc_preproc_prop:
  "LI_preproc_prop (LI_preproc S)"
by (induct S rule: LI_preproc.induct) (auto simp add: LI_preproc_prop_def)

lemma LI_preproc_sem_eq:
  "M; Sc   M; LI_preproc Sc " (is "?A  ?B")
proof
  show "?A  ?B"
  proof (induction S rule: strand_sem_induct)
    case (ConsSnd M ts S)
    hence "M; LI_preproc Sc " "M; map Send1 tsc " using strand_sem_Send_map(5) by auto
    moreover have "ikst (map Send1 ts) set  = {}" unfolding ikst_is_rcv_set by fastforce
    ultimately show ?case using strand_sem_append(1) by simp
  next
    case (ConsRcv M ts S)
    hence "(set ts set )  M; LI_preproc Sc " "M; map Receive1 tsc "
      using strand_sem_Receive_map(3) by auto
    moreover have "ikst (map Receive1 ts) set  = set ts set " unfolding ikst_is_rcv_set by force
    ultimately show ?case using strand_sem_append(1) by (simp add: Un_commute)
  qed simp_all

  show "?B  ?A"
  proof (induction S arbitrary: M rule: LI_preproc.induct)
    case (2 ts S)
    have "ikst (map Send1 ts) set  = {}" unfolding ikst_is_rcv_set by fastforce
    hence "M; Sc " "M; map Send1 tsc " using 2 strand_sem_append(1) by auto
    thus ?case using strand_sem_Send_map(5) by simp
  next
    case (3 ts S)
    have "ikst (map Receive1 ts) set  = set ts set " unfolding ikst_is_rcv_set by force
    hence "M  (set ts set ); Sc " "M; map Receive1 tsc "
      using 3 strand_sem_append(1) by auto
    thus ?case using strand_sem_Receive_map(3) by (simp add: Un_commute)
  qed simp_all
qed

lemma LI_preproc_sem_eq':
  "( c S, θ)  ( c LI_preproc S, θ)"
using LI_preproc_sem_eq unfolding constr_sem_c_def by simp

lemma LI_preproc_vars_eq:
  "fvst (LI_preproc S) = fvst S"
  "bvarsst (LI_preproc S) = bvarsst S"
  "varsst (LI_preproc S) = varsst S"
by (induct S rule: LI_preproc.induct) auto

lemma LI_preproc_trms_eq:
  "trmsst (LI_preproc S) = trmsst S"
by (induct S rule: LI_preproc.induct) auto

lemma LI_preproc_wfst:
  assumes "wfst X S"
  shows "wfst X (LI_preproc S)"
  using assms
proof (induction S arbitrary: X rule: wfst_induct)
  case (ConsRcv X ts S)
  hence "fvset (set ts)  X" "wfst X (LI_preproc S)" by auto
  thus ?case using wf_Receive1_prefix by simp
next
  case (ConsSnd X ts S)
  hence "wfst (X  fvset (set ts)) (LI_preproc S)" by simp
  thus ?case using wf_Send1_prefix by simp
qed simp_all

lemma LI_preproc_preserves_wellformedness:
  assumes "wfconstr S θ"
  shows "wfconstr (LI_preproc S) θ"
using assms LI_preproc_vars_eq[of S] LI_preproc_wfst[of "{}" S] unfolding wfconstr_def by argo

lemma LI_preproc_prop_SendE:
  assumes "LI_preproc_prop S"
    and "Send ts  set S"
  shows "(x. ts = [Var x])  (f T. ts = [Fun f T])"
proof -
  obtain t where "ts = [t]" using assms unfolding LI_preproc_prop_def by auto
  thus ?thesis by (cases t) auto
qed

lemma LI_preproc_prop_split:
  "LI_preproc_prop (S@S')  LI_preproc_prop S  LI_preproc_prop S'" (is "?A  ?B")
proof
  show "?A  ?B"
  proof (induction S)
    case (Cons x S) thus ?case unfolding LI_preproc_prop_def by (cases x) auto
  qed (simp add: LI_preproc_prop_def)

  show "?B  ?A"
  proof (induction S)
    case (Cons x S) thus ?case unfolding LI_preproc_prop_def by (cases x) auto
  qed (simp add: LI_preproc_prop_def)
qed

subsection ‹Lemma: The Lazy Intruder is Well-founded›
context
begin
private lemma LI_compose_measure_lt:
  "((S@(map Send1 T)@S',θ1), (S@Send [Fun f T]#S',θ2))  measurest"
using strand_fv_card_map_fun_eq[of S f T S'] strand_size_map_fun_lt[of T f]
by (simp add: measurest_def sizest_def)

private lemma LI_unify_measure_lt:
  assumes "Some δ = mgu (Fun f T) t" "fv t  fvst S"
  shows "(((S@S') st δ,θ1), (S@Send [Fun f T]#S',θ2))  measurest"
proof (cases "δ = Var")
  assume "δ = Var"
  hence "(S@S') st δ = S@S'" by blast
  thus ?thesis
    using strand_fv_card_rm_fun_le[of S S' f T]
    by (auto simp add: measurest_def sizest_def)
next
  assume "δ  Var"
  then obtain v where "v  fv (Fun f T)  fv t" "subst_elim δ v"
    using mgu_eliminates[OF assms(1)[symmetric]] by metis
  hence v_in: "v  fvst (S@Send [Fun f T]#S')"
    using assms(2) by (auto simp add: measurest_def sizest_def)
  
  have "range_vars δ  fv (Fun f T)  fvst S"
    using assms(2) mgu_vars_bounded[OF assms(1)[symmetric]] by auto
  hence img_bound: "range_vars δ  fvst (S@Send [Fun f T]#S')" by auto

  have finite_fv: "finite (fvst (S@Send [Fun f T]#S'))" by auto

  have "v  fvst ((S@Send [Fun f T]#S') st δ)"
    using strand_fv_subst_subset_if_subst_elim[OF subst_elim δ v] v_in by metis
  hence v_not_in: "v  fvst ((S@S') st δ)" by auto
  
  have "fvst ((S@S') st δ)  fvst (S@Send [Fun f T]#S')"
    using strand_subst_fv_bounded_if_img_bounded[OF img_bound] by simp
  hence "fvst ((S@S') st δ)  fvst (S@Send [Fun f T]#S')" using v_in v_not_in by blast
  hence "card (fvst ((S@S') st δ)) < card (fvst (S@Send [Fun f T]#S'))"
    using psubset_card_mono[OF finite_fv] by simp
  thus ?thesis by (auto simp add: measurest_def sizest_def)
qed

private lemma LI_equality_measure_lt:
  assumes "Some δ = mgu t t'"
  shows "(((S@S') st δ,θ1), (S@Equality a t t'#S',θ2))  measurest"
proof (cases "δ = Var")
  assume "δ = Var"
  hence "(S@S') st δ = S@S'" by blast
  thus ?thesis
    using strand_fv_card_rm_eq_le[of S S' a t t']
    by (auto simp add: measurest_def sizest_def)
next
  assume "δ  Var"
  then obtain v where "v  fv t  fv t'" "subst_elim δ v"
    using mgu_eliminates[OF assms(1)[symmetric]] by metis
  hence v_in: "v  fvst (S@Equality a t t'#S')" using assms by auto
  
  have "range_vars δ  fv t  fv t'  fvst S"
    using assms mgu_vars_bounded[OF assms(1)[symmetric]] by auto
  hence img_bound: "range_vars δ  fvst (S@Equality a t t'#S')" by auto

  have finite_fv: "finite (fvst (S@Equality a t t'#S'))" by auto

  have "v  fvst ((S@Equality a t t'#S') st δ)"
    using strand_fv_subst_subset_if_subst_elim[OF subst_elim δ v] v_in by metis
  hence v_not_in: "v  fvst ((S@S') st δ)" by auto
  
  have "fvst ((S@S') st δ)  fvst (S@Equality a t t'#S')"
    using strand_subst_fv_bounded_if_img_bounded[OF img_bound] by simp
  hence "fvst ((S@S') st δ)  fvst (S@Equality a t t'#S')" using v_in v_not_in by blast
  hence "card (fvst ((S@S') st δ)) < card (fvst (S@Equality a t t'#S'))"
    using psubset_card_mono[OF finite_fv] by simp
  thus ?thesis by (auto simp add: measurest_def sizest_def)
qed

private lemma LI_in_measure: "(S1,θ1)  (S2,θ2)  ((S2,θ2),(S1,θ1))  measurest"
proof (induction rule: LI_rel.induct)
  case (Compose S T f S' θ) thus ?case using LI_compose_measure_lt[of S T S'] by metis
next
  case (Unify S f U δ T S' θ)
  hence "fv (Fun f U)  fvst S"
    using fv_snd_rcv_strand_subset(2)[of S] by force
  thus ?case using LI_unify_measure_lt[OF Unify.hyps(3), of S S'] by metis
qed (metis LI_equality_measure_lt)

private lemma LI_in_measure_trans: "(S1,θ1) + (S2,θ2)  ((S2,θ2),(S1,θ1))  measurest"
by (induction rule: trancl.induct, metis surjective_pairing LI_in_measure)
   (metis (no_types, lifting) surjective_pairing LI_in_measure measurest_trans trans_def)

private lemma LI_converse_wellfounded_trans: "wf ((LI_rel+)¯)"
proof -
  have "(LI_rel+)¯  measurest" using LI_in_measure_trans by auto
  thus ?thesis using measurest_wellfounded wf_subset by metis
qed

private lemma LI_acyclic_trans: "acyclic (LI_rel+)"
using wf_acyclic[OF LI_converse_wellfounded_trans] acyclic_converse by metis

private lemma LI_acyclic: "acyclic LI_rel"
using LI_acyclic_trans acyclic_subset by (simp add: acyclic_def)

lemma LI_no_infinite_chain: "¬(f. i. f i + f (Suc i))"
proof -
  have "¬(f. i. (f (Suc i), f i)  (LI_rel+)¯)"
    using wf_iff_no_infinite_down_chain LI_converse_wellfounded_trans by metis
  thus ?thesis by simp
qed

private lemma LI_unify_finite:
  assumes "finite M"
  shows "finite {((S@Send [Fun f T]#S',θ), ((S@S') st δ,θ s δ)) | δ T'. 
                   simple S  Fun f T'  M  Some δ = mgu (Fun f T) (Fun f T')}"
using assms
proof (induction M rule: finite_induct)
  case (insert m M) thus ?case
  proof (cases m)
    case (Fun g U)
    let ?a = "λδ. ((S@Send [Fun f T]#S',θ), ((S@S') st δ,θ s δ))"
    let ?A = "λB. {?a δ | δ T'. simple S  Fun f T'  B  Some δ = mgu (Fun f T) (Fun f T')}"

    have "?A (insert m M) = (?A M)  (?A {m})" by auto
    moreover have "finite (?A {m})"
    proof (cases "δ. Some δ = mgu (Fun f T) (Fun g U)")
      case True
      then obtain δ where δ: "Some δ = mgu (Fun f T) (Fun g U)" by blast
      
      have A_m_eq: "δ'. ?a δ'  ?A {m}  ?a δ = ?a δ'"
      proof -
        fix δ' assume "?a δ'  ?A {m}"
        hence "σ. Some σ = mgu (Fun f T) (Fun g U)  ?a σ = ?a δ'"
          using m = Fun g U by auto
        thus "?a δ = ?a δ'" by (metis δ option.inject)
      qed

      have "?A {m} = {}  ?A {m} = {?a δ}"
      proof (cases "simple S  ?A {m}  {}")
        case True
        hence "simple S" "?A {m}  {}" by meson+
        hence "?A {m} = {?a δ | δ. Some δ = mgu (Fun f T) (Fun g U)}" using m = Fun g U by auto
        hence "?a δ  ?A {m}" using δ by auto
       show ?thesis
        proof (rule ccontr)
          assume "¬(?A {m} = {}  ?A {m} = {?a δ})"
          then obtain B where B: "?A {m} = insert (?a δ) B" "?a δ  B" "B  {}"
            using ?A {m}  {} ?a δ  ?A {m} by (metis (no_types, lifting) Set.set_insert)
          then obtain b where b: "?a δ  b" "b  B" by (metis (no_types, lifting) ex_in_conv)
          then obtain δ' where δ': "b = ?a δ'" using B(1) by blast
          moreover have "?a δ'  ?A {m}" using B(1) b(2) δ' by auto
          hence "?a δ = ?a δ'" by (blast dest!: A_m_eq)
          ultimately show False using b(1) by simp
        qed
      qed auto
      thus ?thesis by (metis (no_types, lifting) finite.emptyI finite_insert) 
    next
      case False
      hence "?A {m} = {}" using m = Fun g U by blast
      thus ?thesis by (metis finite.emptyI)
    qed
    ultimately show ?thesis using insert.IH by auto
  qed simp
qed fastforce
end


subsection ‹Lemma: The Lazy Intruder Preserves Well-formedness›
context
begin
private lemma LI_preserves_subst_wf_single:
  assumes "(S1,θ1)  (S2,θ2)" "fvst S1  bvarsst S1 = {}" "wfsubst θ1"
  and "subst_domain θ1  varsst S1 = {}" "range_vars θ1  bvarsst S1 = {}"
  shows "fvst S2  bvarsst S2 = {}" "wfsubst θ2"
  and "subst_domain θ2  varsst S2 = {}" "range_vars θ2  bvarsst S2 = {}"
using assms
proof (induction rule: LI_rel.induct)
  case (Compose S X f S' θ)
  { case 1 thus ?case using vars_st_snd_map by auto }
  { case 2 thus ?case using vars_st_snd_map by auto }
  { case 3 thus ?case using vars_st_snd_map by force }
  { case 4 thus ?case using vars_st_snd_map by auto }
next
  case (Unify S f U δ T S' θ)
  hence "fv (Fun f U)  fvst S" using fv_subset_if_in_strand_ik' by blast
  hence *: "subst_domain δ  range_vars δ  fvst (S@Send [Fun f T]#S')"
    using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]]
    unfolding range_vars_alt_def by (fastforce simp del: subst_range.simps)

  have "fvst (S@S')  fvst (S@Send [Fun f T]#S')" "varsst (S@S')  varsst (S@Send [Fun f T]#S')"
    by auto
  hence **: "fvst (S@S' st δ)  fvst (S@Send [Fun f T]#S')"
            "varsst (S@S' st δ)  varsst (S@Send [Fun f T]#S')"
    using subst_sends_strand_fv_to_img[of "S@S'" δ]
          strand_subst_vars_union_bound[of "S@S'" δ] *
    by blast+

  have "wfsubst δ" by (fact mgu_gives_wellformed_subst[OF Unify.hyps(3)[symmetric]])
  
  { case 1
    have "bvarsst (S@S' st δ) = bvarsst (S@Send [Fun f T]#S')"
      using bvars_subst_ident[of "S@S'" δ] by auto
    thus ?case using 1 ** by blast
  }
  { case 2
    hence "subst_domain θ  subst_domain δ = {}" "subst_domain θ  range_vars δ = {}"
      using * by blast+
    thus ?case by (metis wf_subst_compose[OF wfsubst θ wfsubst δ])
  }
  { case 3
    hence "subst_domain θ  varsst (S@S' st δ) = {}" using ** by blast
    moreover have "v  fvst (S@Send [Fun f T]#S')" when "v  subst_domain δ" for v
      using * that by blast
    hence "subst_domain δ  fvst (S@S' st δ) = {}"
      using mgu_eliminates_dom[OF Unify.hyps(3)[symmetric],
                THEN strand_fv_subst_subset_if_subst_elim, of _ "S@Send [Fun f T]#S'"]
      unfolding subst_elim_def by auto
    moreover have "bvarsst (S@S' st δ) = bvarsst (S@Send [Fun f T]#S')"
      using bvars_subst_ident[of "S@S'" δ] by auto
    hence "subst_domain δ  bvarsst (S@S' st δ) = {}" using 3(1) * by blast
    ultimately show ?case
      using ** * subst_domain_compose[of θ δ] varsst_is_fvst_bvarsst[of "S@S' st δ"]
      by blast
  }
  { case 4
    have ***: "bvarsst (S@S' st δ) = bvarsst (S@Send [Fun f T]#S')"
      using bvars_subst_ident[of "S@S'" δ] by auto
    hence "range_vars δ  bvarsst (S@S' st δ) = {}" using 4(1) * by blast
    thus ?case using subst_img_comp_subset[of θ δ] 4(4) *** by blast
  }
next
  case (Equality S δ t t' a S' θ)
  hence *: "subst_domain δ  range_vars δ  fvst (S@Equality a t t'#S')"
    using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]]
    unfolding range_vars_alt_def by fastforce

  have "fvst (S@S')  fvst (S@Equality a t t'#S')" "varsst (S@S')  varsst (S@Equality a t t'#S')"
    by auto
  hence **: "fvst (S@S' st δ)  fvst (S@Equality a t t'#S')"
            "varsst (S@S' st δ)  varsst (S@Equality a t t'#S')"
    using subst_sends_strand_fv_to_img[of "S@S'" δ]
          strand_subst_vars_union_bound[of "S@S'" δ] *
    by blast+

  have "wfsubst δ" by (fact mgu_gives_wellformed_subst[OF Equality.hyps(2)[symmetric]])
  
  { case 1
    have "bvarsst (S@S' st δ) = bvarsst (S@Equality a t t'#S')"
      using bvars_subst_ident[of "S@S'" δ] by auto
    thus ?case using 1 ** by blast
  }
  { case 2
    hence "subst_domain θ  subst_domain δ = {}" "subst_domain θ  range_vars δ = {}"
      using * by blast+
    thus ?case by (metis wf_subst_compose[OF wfsubst θ wfsubst δ])
  }
  { case 3
    hence "subst_domain θ  varsst (S@S' st δ) = {}" using ** by blast
    moreover have "v  fvst (S@Equality a t t'#S')" when "v  subst_domain δ" for v
      using * that by blast
    hence "subst_domain δ  fvst (S@S' st δ) = {}"
      using mgu_eliminates_dom[OF Equality.hyps(2)[symmetric],
                THEN strand_fv_subst_subset_if_subst_elim, of _ "S@Equality a t t'#S'"]
      unfolding subst_elim_def by auto
    moreover have "bvarsst (S@S' st δ) = bvarsst (S@Equality a t t'#S')"
      using bvars_subst_ident[of "S@S'" δ] by auto
    hence "subst_domain δ  bvarsst (S@S' st δ) = {}" using 3(1) * by blast
    ultimately show ?case
      using ** * subst_domain_compose[of θ δ] varsst_is_fvst_bvarsst[of "S@S' st δ"]
      by blast
  }
  { case 4
    have ***: "bvarsst (S@S' st δ) = bvarsst (S@Equality a t t'#S')"
      using bvars_subst_ident[of "S@S'" δ] by auto
    hence "range_vars δ  bvarsst (S@S' st δ) = {}" using 4(1) * by blast
    thus ?case using subst_img_comp_subset[of θ δ] 4(4) *** by blast
  }
qed

private lemma LI_preserves_subst_wf:
  assumes "(S1,θ1) * (S2,θ2)" "fvst S1  bvarsst S1 = {}" "wfsubst θ1"
  and "subst_domain θ1  varsst S1 = {}" "range_vars θ1  bvarsst S1 = {}"
  shows "fvst S2  bvarsst S2 = {}" "wfsubst θ2"
  and "subst_domain θ2  varsst S2 = {}" "range_vars θ2  bvarsst S2 = {}"
using assms
proof (induction S2 θ2 rule: rtrancl_induct2)
  case (step Si θi Sj θj)
  { case 1 thus ?case using LI_preserves_subst_wf_single[OF (Si,θi)  (Sj,θj)] step.IH by metis }
  { case 2 thus ?case using LI_preserves_subst_wf_single[OF (Si,θi)  (Sj,θj)] step.IH by metis }
  { case 3 thus ?case using LI_preserves_subst_wf_single[OF (Si,θi)  (Sj,θj)] step.IH by metis }
  { case 4 thus ?case using LI_preserves_subst_wf_single[OF (Si,θi)  (Sj,θj)] step.IH by metis }
qed metis

lemma LI_preserves_wellformedness:
  assumes "(S1,θ1) * (S2,θ2)" "wfconstr S1 θ1"
  shows "wfconstr S2 θ2"
proof -
  have *: "wfst {} Sj"
    when "(Si, θi)  (Sj, θj)" "wfconstr Si θi" for Si θi Sj θj
    using that
  proof (induction rule: LI_rel.induct)
    case (Compose S T f S' θ) thus ?case by (metis wf_send_compose wfconstr_def)
  next
    case (Unify S f U δ T S' θ)
    have "fv (Fun f T)  fv (Fun f U)  fvst (S@Send [Fun f T]#S')" using Unify.hyps(2) by force
    hence "subst_domain δ  range_vars δ  fvst (S@Send [Fun f T]#S')"
      using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] by (metis subset_trans)
    hence "(subst_domain δ  range_vars δ)  bvarsst (S@Send [Fun f T]#S') = {}"
      using Unify.prems unfolding wfconstr_def by blast
    thus ?case
      using wf_unify[OF _ Unify.hyps(2) MGU_is_Unifier[OF mgu_gives_MGU], of "{}",
                     OF _ Unify.hyps(3)[symmetric], of S'] Unify.prems(1)
      by (auto simp add: wfconstr_def)
  next
    case (Equality S δ t t' a S' θ)
    have "fv t  fv t'  fvst (S@Equality a t t'#S')" using Equality.hyps(2) by force
    hence "subst_domain δ  range_vars δ  fvst (S@Equality a t t'#S')"
      using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] by (metis subset_trans)
    hence "(subst_domain δ  range_vars δ)  bvarsst (S@Equality a t t'#S') = {}"
      using Equality.prems unfolding wfconstr_def by blast
    thus ?case
      using wf_equality[OF _ Equality.hyps(2)[symmetric], of "{}" S a S'] Equality.prems(1)
      by (auto simp add: wfconstr_def)
  qed

  show ?thesis using assms
  proof (induction rule: rtrancl_induct2)
    case (step Si θi Sj θj) thus ?case
      using LI_preserves_subst_wf_single[OF (Si,θi)  (Sj,θj)] *[OF (Si,θi)  (Sj,θj)]
      by (metis wfconstr_def)
  qed simp
qed

lemma LI_preserves_trm_wf:
  assumes "(S,θ) * (S',θ')" "wftrms (trmsst S)"
  shows "wftrms (trmsst S')"
proof -
  { fix S θ S' θ'
    assume "(S,θ)  (S',θ')" "wftrms (trmsst S)"
    hence "wftrms (trmsst S')"
    proof (induction rule: LI_rel.induct)
      case (Compose S T f S' θ)
      hence "wftrm (Fun f T)"
        and *: "t  set S  wftrms (trmsstp t)" "t  set S'  wftrms (trmsstp t)" for t
        by auto
      hence "wftrm t" when "t  set T" for t using that unfolding wftrm_def by auto
      hence "wftrms (trmsstp t)" when "t  set (map Send1 T)" for t
        using that unfolding wftrm_def by auto
      thus ?case using * by force
    next
      case (Unify S f U δ T S' θ)
      have "wftrm (Fun f T)" "wftrm (Fun f U)"
        using Unify.prems(1) Unify.hyps(2) wf_trm_subterm[of _ "Fun f U"]
        by (simp, force)
      hence range_wf: "wftrms (subst_range δ)"
        using mgu_wf_trm[OF Unify.hyps(3)[symmetric]] by simp

      { fix s assume "s  set (S@S' st δ)"
        hence "s'  set (S@S'). s = s' stp δ  wftrms (trmsstp s')"
          using Unify.prems(1) by (auto simp add: subst_apply_strand_def)
        moreover {
          fix s' assume s': "s = s' stp δ" "wftrms (trmsstp s')" "s'  set (S@S')"
          from s'(2) have "trmsstp (s' stp δ) = trmsstp s' set (rm_vars (set (bvarsstp s')) δ)"
          proof (induction s')
            case (Inequality X F) thus ?case by (induct F) (auto simp add: subst_apply_pairs_def)
          qed auto
          hence "wftrms (trmsstp s)"
            using wf_trm_subst[OF wf_trms_subst_rm_vars'[OF range_wf]] wftrms (trmsstp s') s'(1)
            by simp
        }
        ultimately have "wftrms (trmsstp s)" by auto
      }
      thus ?case by auto
    next
      case (Equality S δ t t' a S' θ)
      hence "wftrm t" "wftrm t'" by simp_all
      hence range_wf: "wftrms (subst_range δ)"
        using mgu_wf_trm[OF Equality.hyps(2)[symmetric]] by simp

      { fix s assume "s  set (S@S' st δ)"
        hence "s'  set (S@S'). s = s' stp δ  wftrms (trmsstp s')"
          using Equality.prems(1) by (auto simp add: subst_apply_strand_def)
        moreover {
          fix s' assume s': "s = s' stp δ" "wftrms (trmsstp s')" "s'  set (S@S')"
          from s'(2) have "trmsstp (s' stp δ) = trmsstp s' set (rm_vars (set (bvarsstp s')) δ)"
          proof (induction s')
            case (Inequality X F) thus ?case by (induct F) (auto simp add: subst_apply_pairs_def)
          qed auto
          hence "wftrms (trmsstp s)"
            using wf_trm_subst[OF wf_trms_subst_rm_vars'[OF range_wf]] wftrms (trmsstp s') s'(1)
            by simp
        }
        ultimately have "wftrms (trmsstp s)" by auto
      }
      thus ?case by auto
    qed
  }
  with assms show ?thesis by (induction rule: rtrancl_induct2) metis+
qed

lemma LI_preproc_prop_subst:
  "LI_preproc_prop S  LI_preproc_prop (S st δ)"
proof (induction S)
  case (Cons x S) thus ?case unfolding LI_preproc_prop_def by (cases x) auto
qed (simp add: LI_preproc_prop_def)

lemma LI_preserves_LI_preproc_prop:
  assumes "(S1,θ1) * (S2,θ2)" "LI_preproc_prop S1"
  shows "LI_preproc_prop S2"
using assms
proof (induction rule: rtrancl_induct2)
  case (step Si θi Sj θj)
  hence "LI_preproc_prop Si" by metis
  with step.hyps(2) show ?case 
  proof (induction rule: LI_rel.induct)
    case (Unify S f T' δ T S' θ) thus ?case
      using LI_preproc_prop_subst LI_preproc_prop_split
      by (metis append.left_neutral append_Cons)
  next
    case (Equality S δ t t' uu S' θ) thus ?case
      using LI_preproc_prop_subst LI_preproc_prop_split
      by (metis append.left_neutral append_Cons)
  qed (auto simp add: LI_preproc_prop_def)
qed simp

end

subsection ‹Theorem: Soundness of the Lazy Intruder›
context
begin
private lemma LI_soundness_single:
  assumes "wfconstr S1 θ1" "(S1,θ1)  (S2,θ2)" " c S2,θ2"
  shows " c S1,θ1"
using assms(2,1,3)
proof (induction rule: LI_rel.induct)
  case (Compose S T f S' θ)
  have "ikst (map Send1 T) set θ = {}" by fastforce
  hence *: "{}; Sc " "ikst S set ; map Send1 Tc " "ikst S set ; S'c "
    using Compose unfolding constr_sem_c_def
    by (force, force, fastforce)

  have "ikst S set  c Fun f T  "
    using *(2) Compose.hyps(2) ComposeC[OF _ Compose.hyps(3), of "map (λx. x  ) T"]
    unfolding subst_compose_def by force
  thus " c S@Send [Fun f T]#S',θ"
    using *(1,3)  c S@map Send1 T@S',θ
    by (auto simp add: constr_sem_c_def)
next
  case (Unify S f U δ T S' θ)
  have "(θ s δ) supports " "{}; S@S' st δc "
    using Unify.prems(2) unfolding constr_sem_c_def by metis+
  then obtain σ where σ: "θ s δ s σ = " unfolding subst_compose_def by auto

  have θfun_id: "Fun f U  θ = Fun f U" "Fun f T  θ = Fun f T"
    using Unify.prems(1) subst_apply_term_ident[of "Fun f U" θ]
          fv_subset_if_in_strand_ik[of "Fun f U" S] Unify.hyps(2)
          fv_snd_rcv_strand_subset(2)[of S]
          strand_vars_split(1)[of S "Send [Fun f T]#S'"]
    unfolding wfconstr_def apply blast
    using Unify.prems(1) subst_apply_term_ident[of "Fun f T" θ]
    unfolding wfconstr_def by fastforce
  hence θδ_disj:
      "subst_domain θ  subst_domain δ = {}"
      "subst_domain θ  range_vars δ = {}"
      "subst_domain θ  range_vars θ = {}" 
    using trm_subst_disj mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] apply (blast,blast)
    using Unify.prems(1) unfolding wfconstr_def wfsubst_def by blast
  hence θδ_support: "θ supports " "δ supports "
    by (simp_all add: subst_support_comp_split[OF (θ s δ) supports ])

  have "fv (Fun f T)  fvst (S@Send [Fun f T]#S')" "fv (Fun f U)  fvst (S@Send [Fun f T]#S')"
    using Unify.hyps(2) by force+
  hence δ_vars_bound: "subst_domain δ  range_vars δ  fvst (S@Send [Fun f T]#S')"
    using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] by blast

  have "ikst S set ; [Send [Fun f T]]c "
  proof -
    from Unify.hyps(2) have "Fun f U    ikst S set " by blast
    hence "Fun f U    ikst S set " by blast
    moreover have "Unifier δ (Fun f T) (Fun f U)"
      by (fact MGU_is_Unifier[OF mgu_gives_MGU[OF Unify.hyps(3)[symmetric]]])
    ultimately have "Fun f T    ikst S set "
      using σ by (metis θfun_id subst_subst_compose) 
    thus ?thesis by simp
  qed

  have "{}; Sc " "ikst S set ; S'c "
  proof -
    have "(S@S' st δ) st θ = S@S' st δ" "(S@S') st θ = S@S'"
    proof -
      have "subst_domain θ  varsst (S@S') = {}"
        using Unify.prems(1) by (auto simp add: wfconstr_def)
      hence "subst_domain θ  varsst (S@S' st δ) = {}"
        using θδ_disj(2) strand_subst_vars_union_bound[of "S@S'" δ] by blast
      thus "(S@S' st δ) st θ = S@S' st δ" "(S@S') st θ = S@S'"
        using strand_subst_comp subst_domain θ  varsst (S@S') = {} by (blast,blast)
    qed
    moreover have "subst_idem δ" by (fact mgu_gives_subst_idem[OF Unify.hyps(3)[symmetric]])
    moreover have
        "(subst_domain θ  range_vars θ)  bvarsst (S@S') = {}"
        "(subst_domain θ  range_vars θ)  bvarsst (S@S' st δ) = {}"
        "(subst_domain δ  range_vars δ)  bvarsst (S@S') = {}"
      using wf_constr_bvars_disj[OF Unify.prems(1)]
            wf_constr_bvars_disj'[OF Unify.prems(1) δ_vars_bound]
      by auto
    ultimately have "{}; S@S'c "
      using {}; S@S' st δc  σ
            strand_sem_subst(1)[of θ "S@S' st δ" "{}" "δ s σ"]
            strand_sem_subst(2)[of θ "S@S'" "{}" "δ s σ"] 
            strand_sem_subst_subst_idem[of δ "S@S'" "{}" σ]
      unfolding constr_sem_c_def
      by (metis subst_compose_assoc)
    thus "{}; Sc " "ikst S set ; S'c " by auto
  qed
  
  show " c S@Send [Fun f T]#S',θ"
    using θδ_support(1) ikst S set ; [Send [Fun f T]]c  {}; Sc  ikst S set ; S'c 
    by (auto simp add: constr_sem_c_def)
next
  case (Equality S δ t t' a S' θ)
  have "(θ s δ) supports " "{}; S@S' st δc "
    using Equality.prems(2) unfolding constr_sem_c_def by metis+
  then obtain σ where σ: "θ s δ s σ = " unfolding subst_compose_def by auto

  have "fv t  varsst (S@Equality a t t'#S')" "fv t'  varsst (S@Equality a t t'#S')"
    by auto
  moreover have "subst_domain θ  varsst (S@Equality a t t'#S') = {}"
    using Equality.prems(1) unfolding wfconstr_def by auto
  ultimately have θfun_id: "t  θ = t" "t'  θ = t'" by auto
  hence θδ_disj:
      "subst_domain θ  subst_domain δ = {}"
      "subst_domain θ  range_vars δ = {}"
      "subst_domain θ  range_vars θ = {}" 
    using trm_subst_disj mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] apply (blast,blast)
    using Equality.prems(1) unfolding wfconstr_def wfsubst_def by blast
  hence θδ_support: "θ supports " "δ supports "
    by (simp_all add: subst_support_comp_split[OF (θ s δ) supports ])

  have "fv t  fvst (S@Equality a t t'#S')" "fv t'  fvst (S@Equality a t t'#S')" by auto
  hence δ_vars_bound: "subst_domain δ  range_vars δ  fvst (S@Equality a t t'#S')"
    using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] by blast

  have "ikst S set ; [Equality a t t']c "
  proof -
    have "t  δ = t'  δ"
      using MGU_is_Unifier[OF mgu_gives_MGU[OF Equality.hyps(2)[symmetric]]]
      by metis
    hence "t  (θ s δ) = t'  (θ s δ)" by (metis θfun_id subst_subst_compose)
    hence "t   = t'  " by (metis σ subst_subst_compose) 
    thus ?thesis by simp
  qed

  have "{}; Sc " "ikst S set ; S'c "
  proof -
    have "(S@S' st δ) st θ = S@S' st δ" "(S@S') st θ = S@S'"
    proof -
      have "subst_domain θ  varsst (S@S') = {}"
        using Equality.prems(1)
        by (fastforce simp add: wfconstr_def simp del: subst_range.simps)
      hence "subst_domain θ  fvst (S@S') = {}" by blast
      hence "subst_domain θ  fvst (S@S' st δ) = {}"
        using θδ_disj(2) subst_sends_strand_fv_to_img[of "S@S'" δ] by blast
      thus "(S@S' st δ) st θ = S@S' st δ" "(S@S') st θ = S@S'"
        using strand_subst_comp subst_domain θ  varsst (S@S') = {} by (blast,blast)
    qed
    moreover have
        "(subst_domain θ  range_vars θ)  bvarsst (S@S') = {}"
        "(subst_domain θ  range_vars θ)  bvarsst (S@S' st δ) = {}"
        "(subst_domain δ  range_vars δ)  bvarsst (S@S') = {}"
      using wf_constr_bvars_disj[OF Equality.prems(1)]
            wf_constr_bvars_disj'[OF Equality.prems(1) δ_vars_bound]
      by auto
    ultimately have "{}; S@S'c "
      using {}; S@S' st δc  σ
            strand_sem_subst(1)[of θ "S@S' st δ" "{}" "δ s σ"]
            strand_sem_subst(2)[of θ "S@S'" "{}" "δ s σ"] 
            strand_sem_subst_subst_idem[of δ "S@S'" "{}" σ]
            mgu_gives_subst_idem[OF Equality.hyps(2)[symmetric]]
      unfolding constr_sem_c_def
      by (metis subst_compose_assoc)
    thus "{}; Sc " "ikst S set ; S'c " by auto
  qed
  
  show " c S@Equality a t t'#S',θ"
    using θδ_support(1) ikst S set ; [Equality a t t']c  {}; Sc  ikst S set ; S'c 
    by (auto simp add: constr_sem_c_def)
qed

theorem LI_soundness:
  assumes "wfconstr S1 θ1" "(LI_preproc S1,θ1) * (S2,θ2)" " c S2, θ2"
  shows " c S1, θ1"
using assms(2,1,3)
proof (induction S2 θ2 rule: rtrancl_induct2)
  case (step Si θi Sj θj) thus ?case
    using LI_preproc_preserves_wellformedness[OF wfconstr S1 θ1]
          LI_preserves_wellformedness[OF (LI_preproc S1, θ1) * (Si, θi)]
          LI_soundness_single[OF _ (Si, θi)  (Sj, θj)  c Sj, θj]
    by metis
qed (metis LI_preproc_sem_eq')
end

subsection ‹Theorem: Completeness of the Lazy Intruder›
context
begin
private lemma LI_completeness_single:
  assumes "wfconstr S1 θ1" " c S1, θ1" "¬simple S1" "LI_preproc_prop S1"
  shows "S2 θ2. (S1,θ1)  (S2,θ2)  ( c S2, θ2)"
using not_simple_elim[OF ¬simple S1]
proof -
  { ― ‹In this case S1 isn't simple because it contains an equality constraint,
        so we can simply proceed with the reduction by computing the MGU for the equation›
    assume "S' S'' a t t'. S1 = S'@Equality a t t'#S''  simple S'"
    then obtain S a t t' S' where S1: "S1 = S@Equality a t t'#S'" "simple S" by moura
    hence *: "wfst {} S" " c S, θ1" "θ1 supports " "t   = t'  "
      using  c S1, θ1 wfconstr S1 θ1 wf_eq_fv[of "{}" S t t' S']
            fv_snd_rcv_strand_subset(5)[of S]
      by (auto simp add: constr_sem_c_def wfconstr_def)

    from * have "Unifier  t t'" by simp
    then obtain δ where δ:
        "Some δ = mgu t t'" "subst_idem δ" "subst_domain δ  range_vars δ  fv t  fv t'"
      using mgu_always_unifies mgu_gives_subst_idem mgu_vars_bounded by metis+
    
    have "δ  "
      using mgu_gives_MGU[OF δ(1)[symmetric]]
      by (metis Unifier  t t')
    hence "δ supports " using subst_support_if_mgt_subst_idem[OF _ δ(2)] by metis
    hence "(θ1 s δ) supports " using subst_support_comp θ1 supports  by metis
    
    have "{}; S@S' st δc "
    proof -
      have "subst_domain δ  range_vars δ  fvst S1" using δ(3) S1(1) by auto
      hence "{}; S1 st δc "
        using subst_idem δ δ    c S1, θ1 strand_sem_subst
              wf_constr_bvars_disj'(1)[OF assms(1)]
        unfolding subst_idem_def constr_sem_c_def
        by (metis (no_types) subst_compose_assoc)
      thus "{}; S@S' st δc " using S1(1) by force
    qed
    moreover have "(S@Equality a t t'#S', θ1)  (S@S' st δ, θ1 s δ)"
      using LI_rel.Equality[OF simple S δ(1)] S1 by metis
    ultimately have ?thesis
      using S1(1) (θ1 s δ) supports 
      by (auto simp add: constr_sem_c_def)
  } moreover {
    ― ‹In this case S1 isn't simple because it contains a deduction constraint for a composed
        term, so we must look at how this composed term is derived under the interpretation ℐ›
    assume "S' S'' ts. S1 = S'@Send ts#S''  (x. ts = [Var x])  simple S'"
    hence "S' S'' f T. S1 = S'@Send [Fun f T]#S''  simple S'"
      using LI_preproc_prop_SendE[OF LI_preproc_prop S1]
      by fastforce
    with assms obtain S f T S' where S1: "S1 = S@Send [Fun f T]#S'" "simple S" by moura
    hence "wfst {} S" " c S, θ1" "θ1 supports "
      using  c S1, θ1 wfconstr S1 θ1
      by (auto simp add: constr_sem_c_def wfconstr_def)
  
    ― ‹Lemma for a common subcase›
    have fun_sat: " c S@(map Send1 T)@S', θ1"
      when T: "t. t  set T  ikst S set  c t  "
    proof -
      have "t. t  set T  ikst S set ; [Send1 t]c " using T by simp
      hence "ikst S set ; map Send1 Tc "
        using  c S1, θ1 strand_sem_Send_map by blast 
      moreover have "ikst (S@[Send1 (Fun f T)]) set  = ikst (S@(map Send1 T)) set " by auto
      hence "ikst (S@(map Send1 T)) set ; S'c "
        using  c S1, θ1 unfolding S1(1) constr_sem_c_def by force
      ultimately show ?thesis
        using  c S, θ1 strand_sem_append(1)[of "{}" S  "map Send1 T"]
              strand_sem_append(1)[of "{}" "S@map Send1 T"  S']
        unfolding constr_sem_c_def by simp
    qed
  
    from S1  c S1, θ1 have "ikst S set  c Fun f T  " by (auto simp add: constr_sem_c_def)
    hence ?thesis
    proof cases
      ― ‹Case 1: ℐ(f(T))› has been derived using the AxiomC› rule.›
      case AxiomC
      hence ex_t: "t. t  ikst S  Fun f T   = t  " by auto
      show ?thesis
      proof (cases "T'. Fun f T'  ikst S  Fun f T    Fun f T'  ")
        ― ‹Case 1.1: f(T)› is equal to a variable in the intruder knowledge under ℐ›.
            Hence there must exists a deduction constraint in the simple prefix of the constraint
            in which this variable occurs/"is sent" for the first time. Since this variable itself
            cannot have been derived from the AxiomC› rule (because it must be equal under the
            interpretation to f(T)›, which is by assumption not in the intruder knowledge under
            ℐ›) it must be the case that we can derive it using the ComposeC› rule. Hence we can
            apply the Compose› rule of the lazy intruder to f(T)›.›
        case True
        have "v. Var v  ikst S  Fun f T   =  v"
        proof -
          obtain t where "t  ikst S" "Fun f T   = t  " using ex_t by moura
          thus ?thesis
            using T'. Fun f T'  ikst S  Fun f T    Fun f T'  
            by (cases t) auto
        qed
        hence "v  wfrestrictedvarsst S. Fun f T   =  v"
          using vars_subset_if_in_strand_ik2[of _ S] by fastforce
        then obtain v Spre Ssuf
          where S: "S = Spre@Send [Var v]#Ssuf" "Fun f T   =  v"
                   "¬(w  wfrestrictedvarsst Spre. Fun f T   =  w)"
          using wfst {} S wf_simple_strand_first_Send_var_split[OF _ simple S, of "Fun f T" ]
          by auto
        hence "w. Var w  ikst Spre   v  Var w  " by force
        moreover have "T'. Fun f T'  ikst Spre  Fun f T    Fun f T'  "
          using T'. Fun f T'  ikst S  Fun f T    Fun f T'   S(1)
          by (meson contra_subsetD ik_append_subset(1))
        hence "g T'. Fun g T'  ikst Spre   v  Fun g T'  " using S(2) by simp
        ultimately have "t  ikst Spre.  v  t  " by (metis term.exhaust)
        hence " v  (ikst Spre) set " by auto
  
        have "ikst Spre set  c  v"
          using S1(1) S(1)  c S1, θ1
          by (auto simp add: constr_sem_c_def)
        hence "ikst Spre set  c Fun f T  " using Fun f T   =  v by metis
        hence "length T = arity f" "public f" "t. t  set T  ikst Spre set  c t  "
          using Fun f T   =  v  v  ikst Spre set 
                intruder_synth.simps[of "ikst Spre set " " v"]
          by auto
        hence *: "t. t  set T  ikst S set  c t  "
          using S(1) by (auto intro: ideduct_synth_mono)
        hence " c S@(map Send1 T)@S', θ1" by (metis fun_sat)
        moreover have "(S@Send [Fun f T]#S', θ1)  (S@map Send1 T@S', θ1)"
          by (metis LI_rel.Compose[OF simple S length T = arity f public f])
        ultimately show ?thesis using S1 by auto
      next
        ― ‹Case 1.2: ℐ(f(T))› can be derived from an interpreted composed term in the intruder
            knowledge. Use the Unify› rule on this composed term to further reduce the constraint.›
        case False
        then obtain T' where t: "Fun f T'  ikst S" "Fun f T   = Fun f T'  "
          by auto
        hence "fv (Fun f T')  fvst S1"
          using S1(1) fv_subset_if_in_strand_ik'[OF t(1)]
                fv_snd_rcv_strand_subset(2)[of S]
          by auto
        from t have "Unifier  (Fun f T) (Fun f T')" by simp
        then obtain δ where δ:
            "Some δ = mgu (Fun f T) (Fun f T')" "subst_idem δ"
            "subst_domain δ  range_vars δ  fv (Fun f T)  fv (Fun f T')"
          using mgu_always_unifies mgu_gives_subst_idem mgu_vars_bounded by metis+
        
        have "δ  "
          using mgu_gives_MGU[OF δ(1)[symmetric]]
          by (metis Unifier  (Fun f T) (Fun f T'))
        hence "δ supports " using subst_support_if_mgt_subst_idem[OF _ δ(2)] by metis
        hence "(θ1 s δ) supports " using subst_support_comp θ1 supports  by metis
        
        have "{}; S@S' st δc "
        proof -
          have "subst_domain δ  range_vars δ  fvst S1"
            using δ(3) S1(1) fv (Fun f T')  fvst S1
            unfolding range_vars_alt_def by (fastforce simp del: subst_range.simps)
          hence "{}; S1 st δc "
            using subst_idem δ δ    c S1, θ1 strand_sem_subst
                  wf_constr_bvars_disj'(1)[OF assms(1)]
            unfolding subst_idem_def constr_sem_c_def
            by (metis (no_types) subst_compose_assoc)
          thus "{}; S@S' st δc " using S1(1) by force
        qed
        moreover have "(S@Send [Fun f T]#S', θ1)  (S@S' st δ, θ1 s δ)"
          using LI_rel.Unify[OF simple S t(1) δ(1)] S1 by metis
        ultimately show ?thesis
          using S1(1) (θ1 s δ) supports 
          by (auto simp add: constr_sem_c_def)
      qed
    next
      ― ‹Case 2: ℐ(f(T))› has been derived using the ComposeC› rule.
          Simply use the Compose› rule of the lazy intruder to proceed with the reduction.›
      case (ComposeC T' g)
      hence "f = g" "length T = arity f" "public f"
        and "x. x  set T  ikst S set  c x  "
        by auto
      hence " c S@(map Send1 T)@S', θ1" using fun_sat by metis
      moreover have "(S1, θ1)  (S@(map Send1 T)@S', θ1)"
        using S1 LI_rel.Compose[OF simple S length T = arity f public f]
        by metis
      ultimately show ?thesis by metis
    qed
  } moreover have "A B X F. S1 = A@Inequality X F#B  ineq_model  X F"
    using assms(2) by (auto simp add: constr_sem_c_def)
  ultimately show ?thesis using not_simple_elim[OF ¬simple S1] by metis
qed

theorem LI_completeness:
  assumes "wfconstr S1 θ1" " c S1, θ1"
  shows "S2 θ2. (LI_preproc S1,θ1) * (S2,θ2)  simple S2  ( c S2, θ2)"
proof (cases "simple (LI_preproc S1)")
  case False
  let ?Stuck = "λS2 θ2. ¬(S3 θ3. (S2,θ2)  (S3,θ3)  ( c S3, θ3))"
  let ?Sats = "{((S,θ),(S',θ')). (S,θ)  (S',θ')  ( c S, θ)  ( c S', θ')}"

  have simple_if_stuck:
      "S2 θ2. (LI_preproc S1,θ1) + (S2,θ2);  c S2, θ2; ?Stuck S2 θ2  simple S2"
    using LI_preproc_preserves_wellformedness[OF wfconstr S1 θ1]
          LI_preserves_LI_preproc_prop[OF _ LI_preproc_preproc_prop]
          LI_completeness_single[OF LI_preserves_wellformedness]
          trancl_into_rtrancl
    by metis

  have base: "b. ((LI_preproc S1,θ1),b)  ?Sats"
    using LI_preproc_preserves_wellformedness[OF wfconstr S1 θ1]
          LI_completeness_single[OF _ _ False LI_preproc_preproc_prop]
          LI_preproc_sem_eq'  c S1, θ1
    by auto

  have *: "S θ S' θ'. ((S,θ),(S',θ'))  ?Sats+  (S,θ) + (S',θ')  ( c S', θ')"
  proof -
    fix S θ S' θ'
    assume "((S,θ),(S',θ'))  ?Sats+"
    thus "(S,θ) + (S',θ')  ( c S', θ')"
      by (induct rule: trancl_induct2) auto
  qed

  have "S2 θ2. ((LI_preproc S1,θ1),(S2,θ2))  ?Sats+  ?Stuck S2 θ2"
  proof (rule ccontr)
    assume "¬(S2 θ2. ((LI_preproc S1,θ1),(S2,θ2))  ?Sats+  ?Stuck S2 θ2)"
    hence sat_not_stuck: "S2 θ2. ((LI_preproc S1,θ1),(S2,θ2))  ?Sats+  ¬?Stuck S2 θ2" by blast

    have "S θ. ((LI_preproc S1,θ1),(S,θ))  ?Sats+  (b. ((S,θ),b)  ?Sats)"
    proof (intro allI impI)
      fix S θ assume a: "((LI_preproc S1,θ1),(S,θ))  ?Sats+"
      have "b. ((LI_preproc S1,θ1),b)  ?Sats+  c. b  c  ((LI_preproc S1,θ1),c)  ?Sats+"
      proof -
        fix b assume in_sat: "((LI_preproc S1,θ1),b)  ?Sats+"
        hence "c. (b,c)  ?Sats" using * sat_not_stuck by (cases b) blast
        thus "c. b  c  ((LI_preproc S1,θ1),c)  ?Sats+"
          using trancl_into_trancl[OF in_sat] by blast
      qed
      hence "S' θ'. (S,θ)  (S',θ')  ((LI_preproc S1,θ1),(S',θ'))  ?Sats+" using a by auto
      then obtain S' θ' where S'θ': "(S,θ)  (S',θ')" "((LI_preproc S1,θ1),(S',θ'))  ?Sats+" by auto
      hence " c S', θ'" using * by blast
      moreover have "(LI_preproc S1, θ1) + (S,θ)" using a trancl_mono by blast
      ultimately have "((S,θ),(S',θ'))  ?Sats" using S'θ'(1) * a by blast
      thus "b. ((S,θ),b)  ?Sats" using S'θ'(2) by blast 
    qed
    hence "f. i::nat. (f i, f (Suc i))  ?Sats"
      using infinite_chain_intro'[OF base] by blast
    moreover have "?Sats  LI_rel+" by auto
    hence "¬(f. i::nat. (f i, f (Suc i))  ?Sats)"
      using LI_no_infinite_chain infinite_chain_mono by blast
    ultimately show False by auto
  qed
  hence "S2 θ2. (LI_preproc S1, θ1) + (S2, θ2)  simple S2  ( c S2, θ2)"
    using simple_if_stuck * by blast
  thus ?thesis by (meson trancl_into_rtrancl)
qed (use LI_preproc_sem_eq'  c S1, θ1 in blast) 
end


subsection ‹Corollary: Soundness and Completeness as a Single Theorem›
corollary LI_soundness_and_completeness:
  assumes "wfconstr S1 θ1"
  shows " c S1, θ1  (S2 θ2. (LI_preproc S1,θ1) * (S2,θ2)  simple S2  ( c S2, θ2))"
by (metis LI_soundness[OF assms] LI_completeness[OF assms])

end

end