Theory SourceTargetRelation

theory SourceTargetRelation
  imports Encodings SimulationRelations
begin

section ‹Relation between Source and Target Terms›

subsection ‹Relations Induced by the Encoding Function›

text ‹We map encodability criteria on conditions of relations between source and target terms.
        The encoding function itself induces such relations. To analyse the preservation of source
        term behaviours we use relations that contain the pairs (S, enc S) for all source terms S.
›

inductive_set (in encoding) indRelR
    :: "((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
  where
  encR: "(SourceTerm S, TargetTerm (S))  indRelR"

abbreviation (in encoding) indRelRinfix ::
    "('procS, 'procT) Proc  ('procS, 'procT) Proc  bool" ("_ ℛ⟦⋅⟧R _" [75, 75] 80)
  where
  "P ℛ⟦⋅⟧R Q  (P, Q)  indRelR"

inductive_set (in encoding) indRelRPO
    :: "((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
  where
  encR:   "(SourceTerm S, TargetTerm (S))  indRelRPO" |
  source: "(SourceTerm S, SourceTerm S)  indRelRPO" |
  target: "(TargetTerm T, TargetTerm T)  indRelRPO" |
  trans:  "(P, Q)  indRelRPO; (Q, R)  indRelRPO  (P, R)  indRelRPO"

abbreviation (in encoding) indRelRPOinfix ::
    "('procS, 'procT) Proc  ('procS, 'procT) Proc  bool" ("_ ≲⟦⋅⟧R _" [75, 75] 80)
  where
  "P ≲⟦⋅⟧R Q  (P, Q)  indRelRPO"

lemma (in encoding) indRelRPO_refl:
  shows "refl indRelRPO"
    unfolding refl_on_def
proof auto
  fix P
  show "P ≲⟦⋅⟧R P"
  proof (cases P)
    case (SourceTerm SP)
    assume "SP ∈S P"
    thus "P ≲⟦⋅⟧R P"
      by (simp add: indRelRPO.source)
  next
    case (TargetTerm TP)
    assume "TP ∈T P"
    thus "P ≲⟦⋅⟧R P"
      by (simp add: indRelRPO.target)
  qed
qed

lemma (in encoding) indRelRPO_is_preorder:
  shows "preorder indRelRPO"
    unfolding preorder_on_def
proof
  show "refl indRelRPO"
    by (rule indRelRPO_refl)
next
  show "trans indRelRPO"
    unfolding trans_def
  proof clarify
    fix P Q R
    assume "P ≲⟦⋅⟧R Q" and "Q ≲⟦⋅⟧R R"
    thus "P ≲⟦⋅⟧R R"
      by (rule indRelRPO.trans)
  qed
qed

lemma (in encoding) refl_trans_closure_of_indRelR:
  shows "indRelRPO = indRelR*"
proof auto
  fix P Q
  assume "P ≲⟦⋅⟧R Q"
  thus "(P, Q)  indRelR*"
  proof induct
    case (encR S)
    show "(SourceTerm S, TargetTerm (S))  indRelR*"
        using indRelR.encR[of S]
      by simp
  next
    case (source S)
    show "(SourceTerm S, SourceTerm S)  indRelR*"
      by simp
  next
    case (target T)
    show "(TargetTerm T, TargetTerm T)  indRelR*"
      by simp
  next
    case (trans P Q R)
    assume "(P, Q)  indRelR*" and "(Q, R)  indRelR*"
    thus "(P, R)  indRelR*"
      by simp
  qed
next
  fix P Q
  assume "(P, Q)  indRelR*"
  thus "P ≲⟦⋅⟧R Q"
  proof induct
    show "P ≲⟦⋅⟧R P"
        using indRelRPO_refl
        unfolding refl_on_def
      by simp
  next
    case (step Q R)
    assume "P ≲⟦⋅⟧R Q"
    moreover assume "Q ℛ⟦⋅⟧R R"
    hence "Q ≲⟦⋅⟧R R"
      by (induct, simp add: indRelRPO.encR)
    ultimately show "P ≲⟦⋅⟧R R"
      by (rule indRelRPO.trans)
  qed
qed

text ‹The relation indRelR is the smallest relation that relates all source terms and their
        literal translations. Thus there exists a relation that relates source terms and their
        literal translations and satisfies some predicate on its pairs iff the predicate holds for
        the pairs of indRelR.›

lemma (in encoding) indRelR_impl_exists_source_target_relation:
  fixes PredA :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set  bool"
    and PredB :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc)  bool"
  shows "PredA indRelR  Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  PredA Rel"
    and "(P, Q)  indRelR. PredB (P, Q)
          Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  ((P, Q)  Rel. PredB (P, Q))"
proof -
  have A: "S. SourceTerm S ℛ⟦⋅⟧R TargetTerm (S)"
    by (simp add: indRelR.encR)
  thus "PredA indRelR  Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  PredA Rel"
    by blast
  with A show "(P, Q)  indRelR. PredB (P, Q)
    Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  ((P, Q)  Rel. PredB (P, Q))"
    by blast
qed

lemma (in encoding) source_target_relation_impl_indRelR:
  fixes Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
    and Pred :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc)  bool"
  assumes encRRel: "S. (SourceTerm S, TargetTerm (S))  Rel"
      and condRel: "(P, Q)  Rel. Pred (P, Q)"
  shows "(P, Q)  indRelR. Pred (P, Q)"
proof clarify
  fix P Q
  assume "P ℛ⟦⋅⟧R Q"
  with encRRel have "(P, Q)  Rel"
    by (auto simp add: indRelR.simps)
  with condRel show "Pred (P, Q)"
    by simp
qed

lemma (in encoding) indRelR_iff_exists_source_target_relation:
  fixes Pred :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc)  bool"
  shows "((P, Q)  indRelR. Pred (P, Q))
         = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  ((P, Q)  Rel. Pred (P, Q)))"
      using indRelR_impl_exists_source_target_relation(2)[where PredB="Pred"]
            source_target_relation_impl_indRelR[where Pred="Pred"]
    by blast

lemma (in encoding) indRelR_modulo_pred_impl_indRelRPO_modulo_pred:
  fixes Pred :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc)  bool"
  assumes reflCond:  "P. Pred (P, P)"
      and transCond: "P Q R. Pred (P, Q)  Pred (Q, R)  Pred (P, R)"
  shows "((P, Q)  indRelR. Pred (P, Q)) = ((P, Q)  indRelRPO. Pred (P, Q))"
proof auto
  fix P Q
  assume A: "x  indRelR. Pred x"
  assume "P ≲⟦⋅⟧R Q"
  thus "Pred (P, Q)"
  proof induct
    case (encR S)
    have "SourceTerm S ℛ⟦⋅⟧R TargetTerm (S)"
      by (simp add: indRelR.encR)
    with A show "Pred (SourceTerm S, TargetTerm (S))"
      by simp
  next
    case (source S)
    from reflCond show "Pred (SourceTerm S, SourceTerm S)"
      by simp
  next
    case (target T)
    from reflCond show "Pred (TargetTerm T, TargetTerm T)"
      by simp
  next
    case (trans P Q R)
    assume "Pred (P, Q)" and "Pred (Q, R)"
    with transCond show "Pred (P, R)"
      by blast
  qed
next
  fix P Q
  assume "x  indRelRPO. Pred x" and "P ℛ⟦⋅⟧R Q"
  thus "Pred (P, Q)"
    by (auto simp add: indRelRPO.encR indRelR.simps)
qed

lemma (in encoding) indRelRPO_iff_exists_source_target_relation:
  fixes Pred :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc)  bool"
  shows "((P, Q)  indRelRPO. Pred (P, Q)) = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
          ((P, Q)  Rel. Pred (P, Q))  preorder Rel)"
proof (rule iffI)
  have "S. SourceTerm S ≲⟦⋅⟧R TargetTerm (S)"
    by (simp add: indRelRPO.encR)
  moreover have "preorder indRelRPO"
      using indRelRPO_is_preorder
    by blast
  moreover assume "(P, Q)  indRelRPO. Pred (P, Q)"
  ultimately show "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
                    ((P, Q)  Rel. Pred (P, Q))  preorder Rel"
    by blast
next
  assume "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
           ((P, Q)  Rel. Pred (P, Q))  preorder Rel"
  from this obtain Rel where A1: "S. (SourceTerm S, TargetTerm (S))  Rel"
                         and A2: "(P, Q)  Rel. Pred (P, Q)" and A3: "preorder Rel"
    by blast
  show "(P, Q)  indRelRPO. Pred (P, Q)"
  proof clarify
    fix P Q
    assume "P ≲⟦⋅⟧R Q"
    hence "(P, Q)  Rel"
    proof induct
      case (encR S)
      from A1 show "(SourceTerm S, TargetTerm (S))  Rel"
        by simp
    next
      case (source S)
      from A3 show "(SourceTerm S, SourceTerm S)  Rel"
          unfolding preorder_on_def refl_on_def
        by simp
    next
      case (target T)
      from A3 show "(TargetTerm T, TargetTerm T)  Rel"
          unfolding preorder_on_def refl_on_def
        by simp
    next
      case (trans P Q R)
      assume "(P, Q)  Rel" and "(Q, R)  Rel"
      with A3 show "(P, R)  Rel"
          unfolding preorder_on_def trans_def
        by blast
    qed
    with A2 show "Pred (P, Q)"
      by simp
  qed
qed

text ‹An encoding preserves, reflects, or respects a predicate iff indRelR preserves, reflects,
        or respects this predicate.›

lemma (in encoding) enc_satisfies_pred_impl_indRelR_satisfies_pred:
  fixes Pred :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc)  bool"
  assumes encCond: "S. Pred (SourceTerm S, TargetTerm (S))"
  shows "(P, Q)  indRelR. Pred (P, Q)"
    by (auto simp add: encCond indRelR.simps)

lemma (in encoding) indRelR_satisfies_pred_impl_enc_satisfies_pred:
  fixes Pred :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc)  bool"
  assumes relCond: "(P, Q)  indRelR. Pred (P, Q)"
  shows "S. Pred (SourceTerm S, TargetTerm (S))"
      using relCond indRelR.encR
    by simp

lemma (in encoding) enc_satisfies_pred_iff_indRelR_satisfies_pred:
  fixes Pred :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc)  bool"
  shows "(S. Pred (SourceTerm S, TargetTerm (S))) = ((P, Q)  indRelR. Pred (P, Q))"
      using enc_satisfies_pred_impl_indRelR_satisfies_pred[where Pred="Pred"]
            indRelR_satisfies_pred_impl_enc_satisfies_pred[where Pred="Pred"]
    by blast

lemma (in encoding) enc_satisfies_binary_pred_iff_indRelR_satisfies_binary_pred:
  fixes Pred :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc)  'b  bool"
  shows "(S a. Pred (SourceTerm S, TargetTerm (S)) a) = ((P, Q)  indRelR. a. Pred (P, Q) a)"
      using enc_satisfies_pred_iff_indRelR_satisfies_pred
    by simp

lemma (in encoding) enc_preserves_pred_iff_indRelR_preserves_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_preserves_pred Pred = rel_preserves_pred indRelR Pred"
      using enc_satisfies_pred_iff_indRelR_satisfies_pred[where Pred="λ(P, Q). Pred P  Pred Q"]
    by blast

lemma (in encoding) enc_preserves_binary_pred_iff_indRelR_preserves_binary_pred:
  fixes Pred :: "('procS, 'procT) Proc  'b  bool"
  shows "enc_preserves_binary_pred Pred = rel_preserves_binary_pred indRelR Pred"
      using enc_satisfies_binary_pred_iff_indRelR_satisfies_binary_pred[where
             Pred="λ(P, Q) a. Pred P a  Pred Q a"]
    by blast

lemma (in encoding) enc_preserves_pred_iff_indRelRPO_preserves_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_preserves_pred Pred = rel_preserves_pred indRelRPO Pred"
      using enc_preserves_pred_iff_indRelR_preserves_pred[where Pred="Pred"]
            indRelR_modulo_pred_impl_indRelRPO_modulo_pred[where
             Pred="λ(P, Q). Pred P  Pred Q"]
    by blast

lemma (in encoding) enc_reflects_pred_iff_indRelR_reflects_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_reflects_pred Pred = rel_reflects_pred indRelR Pred"
      using enc_satisfies_pred_iff_indRelR_satisfies_pred[where Pred="λ(P, Q). Pred Q  Pred P"]
    by blast

lemma (in encoding) enc_reflects_binary_pred_iff_indRelR_reflects_binary_pred:
  fixes Pred :: "('procS, 'procT) Proc  'b  bool"
  shows "enc_reflects_binary_pred Pred = rel_reflects_binary_pred indRelR Pred"
      using enc_satisfies_binary_pred_iff_indRelR_satisfies_binary_pred[where
             Pred="λ(P, Q) a. Pred Q a  Pred P a"]
    by blast

lemma (in encoding) enc_reflects_pred_iff_indRelRPO_reflects_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_reflects_pred Pred = rel_reflects_pred indRelRPO Pred"
      using enc_reflects_pred_iff_indRelR_reflects_pred[where Pred="Pred"]
            indRelR_modulo_pred_impl_indRelRPO_modulo_pred[where
             Pred="λ(P, Q). Pred Q  Pred P"]
    by blast

lemma (in encoding) enc_respects_pred_iff_indRelR_respects_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_respects_pred Pred = rel_respects_pred indRelR Pred"
      using enc_preserves_pred_iff_indRelR_preserves_pred[where Pred="Pred"]
            enc_reflects_pred_iff_indRelR_reflects_pred[where Pred="Pred"]
    by blast

lemma (in encoding) enc_respects_binary_pred_iff_indRelR_respects_binary_pred:
  fixes Pred :: "('procS, 'procT) Proc  'b  bool"
  shows "enc_respects_binary_pred Pred = rel_respects_binary_pred indRelR Pred"
      using enc_preserves_binary_pred_iff_indRelR_preserves_binary_pred[where Pred="Pred"]
            enc_reflects_binary_pred_iff_indRelR_reflects_binary_pred[where Pred="Pred"]
    by blast

lemma (in encoding) enc_respects_pred_iff_indRelRPO_respects_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_respects_pred Pred = rel_respects_pred indRelRPO Pred"
      using enc_respects_pred_iff_indRelR_respects_pred[where Pred="Pred"]
            indRelR_modulo_pred_impl_indRelRPO_modulo_pred[where Pred="λ(P, Q). Pred Q = Pred P"]
    apply simp by blast

text ‹Accordingly an encoding preserves, reflects, or respects a predicate iff there exists a
        relation that relates source terms with their literal translations and preserves, reflects,
        or respects this predicate.›

lemma (in encoding) enc_satisfies_pred_iff_source_target_satisfies_pred:
  fixes Pred :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc)  bool"
  shows "(S. Pred (SourceTerm S, TargetTerm (S)))
         = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  ((P, Q)  Rel. Pred (P, Q)))"
    and "P Q R. Pred (P, Q)  Pred (Q, R)  Pred (P, R); P. Pred (P, P) 
         (S. Pred (SourceTerm S, TargetTerm (S))) = (Rel. (S.
         (SourceTerm S, TargetTerm (S))  Rel)  ((P, Q)  Rel. Pred (P, Q))  preorder Rel)"
proof -
  show "(S. Pred (SourceTerm S, TargetTerm (S)))
        = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  ((P, Q)  Rel. Pred (P, Q)))"
      using enc_satisfies_pred_iff_indRelR_satisfies_pred[where Pred="Pred"]
            indRelR_iff_exists_source_target_relation[where Pred="Pred"]
    by simp
next
  have "(S. Pred (SourceTerm S, TargetTerm (S))) = ((P, Q)  indRelR. Pred (P, Q))"
      using enc_satisfies_pred_iff_indRelR_satisfies_pred[where Pred="Pred"]
    by simp
  moreover assume "P Q R. Pred (P, Q)  Pred (Q, R)  Pred (P, R)" and "P. Pred (P, P)"
  hence "((P, Q)  indRelR. Pred (P, Q)) = ((P, Q)  indRelRPO. Pred (P, Q))"
      using indRelR_modulo_pred_impl_indRelRPO_modulo_pred[where Pred="Pred"]
    by blast
  ultimately show "(S. Pred (SourceTerm S, TargetTerm (S))) = (Rel.
   (S. (SourceTerm S, TargetTerm (S))  Rel)  ((P, Q)  Rel. Pred (P, Q))  preorder Rel)"
      using indRelRPO_iff_exists_source_target_relation[where Pred="Pred"]
    by simp
qed

lemma (in encoding) enc_preserves_pred_iff_source_target_rel_preserves_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_preserves_pred Pred
         = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_preserves_pred Rel Pred)"
    and "enc_preserves_pred Pred = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
          rel_preserves_pred Rel Pred  preorder Rel)"
proof -
  have A1: "enc_preserves_pred Pred
            = (S. (λ(P, Q). Pred P  Pred Q) (SourceTerm S, TargetTerm (S)))"
    by blast
  moreover have A2: "Rel. rel_preserves_pred Rel Pred
                     = ((P, Q)  Rel. (λ(P, Q). Pred P  Pred Q) (P, Q))"
    by blast
  ultimately show "enc_preserves_pred Pred = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
                    rel_preserves_pred Rel Pred)"
      using enc_satisfies_pred_iff_source_target_satisfies_pred(1)[where
             Pred="λ(P, Q). Pred P  Pred Q"]
    by simp
  from A1 A2 show "enc_preserves_pred Pred = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
                    rel_preserves_pred Rel Pred  preorder Rel)"
      using enc_satisfies_pred_iff_source_target_satisfies_pred(2)[where
             Pred="λ(P, Q). Pred P  Pred Q"]
    by simp
qed

lemma (in encoding) enc_preserves_binary_pred_iff_source_target_rel_preserves_binary_pred:
  fixes Pred :: "('procS, 'procT) Proc  'b  bool"
  shows "enc_preserves_binary_pred Pred = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
          rel_preserves_binary_pred Rel Pred)"
proof -
  have "enc_preserves_binary_pred Pred
        = (S. (λ(P, Q). a. Pred P a  Pred Q a) (SourceTerm S, TargetTerm (S)))"
    by blast
  moreover have "Rel. rel_preserves_binary_pred Rel Pred
                 = ((P, Q)  Rel. (λ(P, Q). a. Pred P a  Pred Q a) (P, Q))"
    by blast
  ultimately show "enc_preserves_binary_pred Pred = (Rel. (S.
                   (SourceTerm S, TargetTerm (S))  Rel)  rel_preserves_binary_pred Rel Pred)"
      using enc_satisfies_pred_iff_source_target_satisfies_pred(1)[where
             Pred="λ(P, Q). a. Pred P a  Pred Q a"]
    by simp
qed

lemma (in encoding) enc_reflects_pred_iff_source_target_rel_reflects_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_reflects_pred Pred
         = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_reflects_pred Rel Pred)"
    and "enc_reflects_pred Pred = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
          rel_reflects_pred Rel Pred  preorder Rel)"
proof -
  have A1: "enc_reflects_pred Pred
        = (S. (λ(P, Q). Pred Q  Pred P) (SourceTerm S, TargetTerm (S)))"
    by blast
  moreover have A2: "Rel. rel_reflects_pred Rel Pred
                     = ((P, Q)  Rel. (λ(P, Q). Pred Q  Pred P) (P, Q))"
    by blast
  ultimately show "enc_reflects_pred Pred = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
                    rel_reflects_pred Rel Pred)"
      using enc_satisfies_pred_iff_source_target_satisfies_pred(1)[where
             Pred="λ(P, Q). Pred Q  Pred P"]
    by simp
  from A1 A2 show "enc_reflects_pred Pred = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
                    rel_reflects_pred Rel Pred  preorder Rel)"
      using enc_satisfies_pred_iff_source_target_satisfies_pred(2)[where
             Pred="λ(P, Q). Pred Q  Pred P"]
    by simp
qed

lemma (in encoding) enc_reflects_binary_pred_iff_source_target_rel_reflects_binary_pred:
  fixes Pred :: "('procS, 'procT) Proc  'b  bool"
  shows "enc_reflects_binary_pred Pred = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
          rel_reflects_binary_pred Rel Pred)"
proof -
  have "enc_reflects_binary_pred Pred
        = (S. (λ(P, Q). a. Pred Q a  Pred P a) (SourceTerm S, TargetTerm (S)))"
    by blast
  moreover have "Rel. rel_reflects_binary_pred Rel Pred
                 = ((P, Q)  Rel. (λ(P, Q). a. Pred Q a  Pred P a) (P, Q))"
    by blast
  ultimately show "enc_reflects_binary_pred Pred = (Rel. (S.
                   (SourceTerm S, TargetTerm (S))  Rel)  rel_reflects_binary_pred Rel Pred)"
      using enc_satisfies_pred_iff_source_target_satisfies_pred(1)[where
             Pred="λ(P, Q). a. Pred Q a  Pred P a"]
    by simp
qed

lemma (in encoding) enc_respects_pred_iff_source_target_rel_respects_pred_encR:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_respects_pred Pred
         = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_respects_pred Rel Pred)"
    and "enc_respects_pred Pred = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
          rel_respects_pred Rel Pred  preorder Rel)"
proof -
  have A1: "enc_respects_pred Pred
            = (S. (λ(P, Q). Pred P = Pred Q) (SourceTerm S, TargetTerm (S)))"
    by blast
  moreover
  have A2: "Rel. rel_respects_pred Rel Pred = ((P, Q)  Rel. (λ(P, Q). Pred P = Pred Q) (P, Q))"
    by blast
  ultimately show "enc_respects_pred Pred = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
                    rel_respects_pred Rel Pred)"
      using enc_satisfies_pred_iff_source_target_satisfies_pred(1)[where
             Pred="λ(P, Q). Pred P = Pred Q"]
    by simp
  from A1 A2 show "enc_respects_pred Pred = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
                    rel_respects_pred Rel Pred  preorder Rel)"
      using enc_satisfies_pred_iff_source_target_satisfies_pred(2)[where
             Pred="λ(P, Q). Pred P = Pred Q"]
    by simp
qed

lemma (in encoding) enc_respects_binary_pred_iff_source_target_rel_respects_binary_pred_encR:
  fixes Pred :: "('procS, 'procT) Proc  'b  bool"
  shows "enc_respects_binary_pred Pred = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
          rel_respects_binary_pred Rel Pred)"
proof -
  have "enc_respects_binary_pred Pred
        = (S. (λ(P, Q). a. Pred P a = Pred Q a) (SourceTerm S, TargetTerm (S)))"
    by blast
  moreover have "Rel. rel_respects_binary_pred Rel Pred
                 = ((P, Q)  Rel. (λ(P, Q). a. Pred P a = Pred Q a) (P, Q))"
    by blast
  ultimately show "enc_respects_binary_pred Pred = (Rel. (S.
                   (SourceTerm S, TargetTerm (S))  Rel)  rel_respects_binary_pred Rel Pred)"
      using enc_satisfies_pred_iff_source_target_satisfies_pred(1)[where
             Pred="λ(P, Q). a. Pred P a = Pred Q a"]
    by simp
qed

text ‹To analyse the reflection of source term behaviours we use relations that contain the pairs
        (enc S, S) for all source terms S.›

inductive_set (in encoding) indRelL
    :: "((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
  where
  encL: "(TargetTerm (S), SourceTerm S)  indRelL"

abbreviation (in encoding) indRelLinfix ::
    "('procS, 'procT) Proc  ('procS, 'procT) Proc  bool" ("_ ℛ⟦⋅⟧L _" [75, 75] 80)
  where
  "P ℛ⟦⋅⟧L Q  (P, Q)  indRelL"

inductive_set (in encoding) indRelLPO
    :: "((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
  where
  encL:   "(TargetTerm (S), SourceTerm S)  indRelLPO" |
  source: "(SourceTerm S, SourceTerm S)  indRelLPO" |
  target: "(TargetTerm T, TargetTerm T)  indRelLPO" |
  trans:  "(P, Q)  indRelLPO; (Q, R)  indRelLPO  (P, R)  indRelLPO"

abbreviation (in encoding) indRelLPOinfix ::
    "('procS, 'procT) Proc  ('procS, 'procT) Proc  bool" ("_ ≲⟦⋅⟧L _" [75, 75] 80)
  where
  "P ≲⟦⋅⟧L Q  (P, Q)  indRelLPO"

lemma (in encoding) indRelLPO_refl:
  shows "refl indRelLPO"
    unfolding refl_on_def
proof auto
  fix P
  show "P ≲⟦⋅⟧L P"
  proof (cases P)
    case (SourceTerm SP)
    assume "SP ∈S P"
    thus "P ≲⟦⋅⟧L P"
      by (simp add: indRelLPO.source)
  next
    case (TargetTerm TP)
    assume "TP ∈T P"
    thus "P ≲⟦⋅⟧L P"
      by (simp add: indRelLPO.target)
  qed
qed

lemma (in encoding) indRelLPO_is_preorder:
  shows "preorder indRelLPO"
    unfolding preorder_on_def
proof
  show "refl indRelLPO"
    by (rule indRelLPO_refl)
next
  show "trans indRelLPO"
    unfolding trans_def
  proof clarify
    fix P Q R
    assume "P ≲⟦⋅⟧L Q" and "Q ≲⟦⋅⟧L R"
    thus "P ≲⟦⋅⟧L R"
      by (rule indRelLPO.trans)
  qed
qed

lemma (in encoding) refl_trans_closure_of_indRelL:
  shows "indRelLPO = indRelL*"
proof auto
  fix P Q
  assume "P ≲⟦⋅⟧L Q"
  thus "(P, Q)  indRelL*"
  proof induct
    case (encL S)
    show "(TargetTerm (S), SourceTerm S)  indRelL*"
        using indRelL.encL[of S]
      by simp
  next
    case (source S)
    show "(SourceTerm S, SourceTerm S)  indRelL*"
      by simp
  next
    case (target T)
    show "(TargetTerm T, TargetTerm T)  indRelL*"
      by simp
  next
    case (trans P Q R)
    assume "(P, Q)  indRelL*" and "(Q, R)  indRelL*"
    thus "(P, R)  indRelL*"
      by simp
  qed
next
  fix P Q
  assume "(P, Q)  indRelL*"
  thus "P ≲⟦⋅⟧L Q"
  proof induct
    show "P ≲⟦⋅⟧L P"
        using indRelLPO_refl
        unfolding refl_on_def
      by simp
  next
    case (step Q R)
    assume "P ≲⟦⋅⟧L Q"
    moreover assume "Q ℛ⟦⋅⟧L R"
    hence "Q ≲⟦⋅⟧L R"
      by (induct, simp add: indRelLPO.encL)
    ultimately show "P ≲⟦⋅⟧L R"
      by (simp add: indRelLPO.trans[of P Q R])
  qed
qed

text ‹The relations indRelR and indRelL are dual. indRelR preserves some predicate iff indRelL
        reflects it. indRelR reflects some predicate iff indRelL reflects it. indRelR respects some
        predicate iff indRelL does.›

lemma (in encoding) indRelR_preserves_pred_iff_indRelL_reflects_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "rel_preserves_pred indRelR Pred = rel_reflects_pred indRelL Pred"
proof
  assume preservation: "rel_preserves_pred indRelR Pred"
  show "rel_reflects_pred indRelL Pred"
  proof clarify
    fix P Q
    assume "P ℛ⟦⋅⟧L Q"
    from this obtain S where "S ∈S Q" and "S ∈T P"
      by (induct, blast)
    hence "Q ℛ⟦⋅⟧R P"
      by (simp add: indRelR.encR)
    moreover assume "Pred Q"
    ultimately show "Pred P"
        using preservation
      by blast
  qed
next
  assume reflection: "rel_reflects_pred indRelL Pred"
  show "rel_preserves_pred indRelR Pred"
  proof clarify
    fix P Q
    assume "P ℛ⟦⋅⟧R Q"
    from this obtain S where "S ∈S P" and "S ∈T Q"
      by (induct, blast)
    hence "Q ℛ⟦⋅⟧L P"
      by (simp add: indRelL.encL)
    moreover assume "Pred P"
    ultimately show "Pred Q"
        using reflection
      by blast
  qed
qed

lemma (in encoding) indRelR_preserves_binary_pred_iff_indRelL_reflects_binary_pred:
  fixes Pred :: "('procS, 'procT) Proc  'b  bool"
  shows "rel_preserves_binary_pred indRelR Pred = rel_reflects_binary_pred indRelL Pred"
proof
  assume preservation: "rel_preserves_binary_pred indRelR Pred"
  show "rel_reflects_binary_pred indRelL Pred"
  proof clarify
    fix P Q x
    assume "P ℛ⟦⋅⟧L Q"
    from this obtain S where "S ∈S Q" and "S ∈T P"
      by (induct, blast)
    hence "Q ℛ⟦⋅⟧R P"
      by (simp add: indRelR.encR)
    moreover assume "Pred Q x"
    ultimately show "Pred P x"
        using preservation
      by blast
  qed
next
  assume reflection: "rel_reflects_binary_pred indRelL Pred"
  show "rel_preserves_binary_pred indRelR Pred"
  proof clarify
    fix P Q x
    assume "P ℛ⟦⋅⟧R Q"
    from this obtain S where "S ∈S P" and "S ∈T Q"
      by (induct, blast)
    hence "Q ℛ⟦⋅⟧L P"
      by (simp add: indRelL.encL)
    moreover assume "Pred P x"
    ultimately show "Pred Q x"
        using reflection
      by blast
  qed
qed

lemma (in encoding) indRelR_reflects_pred_iff_indRelL_preserves_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "rel_reflects_pred indRelR Pred = rel_preserves_pred indRelL Pred"
proof
  assume reflection: "rel_reflects_pred indRelR Pred"
  show "rel_preserves_pred indRelL Pred"
  proof clarify
    fix P Q
    assume "P ℛ⟦⋅⟧L Q"
    from this obtain S where "S ∈S Q" and "S ∈T P"
      by (induct, blast)
    hence "Q ℛ⟦⋅⟧R P"
      by (simp add: indRelR.encR)
    moreover assume "Pred P"
    ultimately show "Pred Q"
        using reflection
      by blast
  qed
next
  assume preservation: "rel_preserves_pred indRelL Pred"
  show "rel_reflects_pred indRelR Pred"
  proof clarify
    fix P Q
    assume "P ℛ⟦⋅⟧R Q"
    from this obtain S where "S ∈S P" and "S ∈T Q"
      by (induct, blast)
    hence "Q ℛ⟦⋅⟧L P"
      by (simp add: indRelL.encL)
    moreover assume "Pred Q"
    ultimately show "Pred P"
        using preservation
      by blast
  qed
qed

lemma (in encoding) indRelR_reflects_binary_pred_iff_indRelL_preserves_binary_pred:
  fixes Pred :: "('procS, 'procT) Proc  'b  bool"
  shows "rel_reflects_binary_pred indRelR Pred = rel_preserves_binary_pred indRelL Pred"
proof
  assume reflection: "rel_reflects_binary_pred indRelR Pred"
  show "rel_preserves_binary_pred indRelL Pred"
  proof clarify
    fix P Q x
    assume "P ℛ⟦⋅⟧L Q"
    from this obtain S where "S ∈S Q" and "S ∈T P"
      by (induct, blast)
    hence "Q ℛ⟦⋅⟧R P"
      by (simp add: indRelR.encR)
    moreover assume "Pred P x"
    ultimately show "Pred Q x"
        using reflection
      by blast
  qed
next
  assume preservation: "rel_preserves_binary_pred indRelL Pred"
  show "rel_reflects_binary_pred indRelR Pred"
  proof clarify
    fix P Q x
    assume "P ℛ⟦⋅⟧R Q"
    from this obtain S where "S ∈S P" and "S ∈T Q"
      by (induct, blast)
    hence "Q ℛ⟦⋅⟧L P"
      by (simp add: indRelL.encL)
    moreover assume "Pred Q x"
    ultimately show "Pred P x"
        using preservation
      by blast
  qed
qed

lemma (in encoding) indRelR_respects_pred_iff_indRelL_respects_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "rel_respects_pred indRelR Pred = rel_respects_pred indRelL Pred"
      using indRelR_preserves_pred_iff_indRelL_reflects_pred[where Pred="Pred"]
            indRelR_reflects_pred_iff_indRelL_preserves_pred[where Pred="Pred"]
    by blast

lemma (in encoding) indRelR_respects_binary_pred_iff_indRelL_respects_binary_pred:
  fixes Pred :: "('procS, 'procT) Proc 'b  bool"
  shows "rel_respects_binary_pred indRelR Pred = rel_respects_binary_pred indRelL Pred"
      using indRelR_preserves_binary_pred_iff_indRelL_reflects_binary_pred[where Pred="Pred"]
            indRelR_reflects_binary_pred_iff_indRelL_preserves_binary_pred[where Pred="Pred"]
    by blast

lemma (in encoding) indRelR_cond_preservation_iff_indRelL_cond_reflection:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "(Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_preserves_pred Rel Pred)
         = (Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_reflects_pred Rel Pred)"
proof
  assume "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_preserves_pred Rel Pred"
  then obtain Rel where A1: "S. (SourceTerm S, TargetTerm (S))  Rel"
                    and A2: "rel_preserves_pred Rel Pred"
    by blast
  from A1 have "S. (TargetTerm (S), SourceTerm S)  Rel¯"
    by simp
  moreover from A2 have "rel_reflects_pred (Rel¯) Pred"
    by simp
  ultimately show "Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_reflects_pred Rel Pred"
    by blast
next
  assume "Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_reflects_pred Rel Pred"
  then obtain Rel where B1: "S. (TargetTerm (S), SourceTerm S)  Rel"
                    and B2: "rel_reflects_pred Rel Pred"
    by blast
  from B1 have "S. (SourceTerm S, TargetTerm (S))  Rel¯"
    by simp
  moreover from B2 have "rel_preserves_pred (Rel¯) Pred"
    by blast
  ultimately
  show "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_preserves_pred Rel Pred"
    by blast
qed

lemma (in encoding) indRelR_cond_binary_preservation_iff_indRelL_cond_binary_reflection:
  fixes Pred :: "('procS, 'procT) Proc  'b  bool"
  shows "(Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_preserves_binary_pred Rel Pred)
         = (Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)
             rel_reflects_binary_pred Rel Pred)"
proof
  assume "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_preserves_binary_pred Rel Pred"
  then obtain Rel where A1: "S. (SourceTerm S, TargetTerm (S))  Rel"
                    and A2: "rel_preserves_binary_pred Rel Pred"
    by blast
  from A1 have "S. (TargetTerm (S), SourceTerm S)  Rel¯"
    by simp
  moreover from A2 have "rel_reflects_binary_pred (Rel¯) Pred"
    by simp
  ultimately
  show "Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_reflects_binary_pred Rel Pred"
    by blast
next
  assume "Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_reflects_binary_pred Rel Pred"
  then obtain Rel where B1: "S. (TargetTerm (S), SourceTerm S)  Rel"
                    and B2: "rel_reflects_binary_pred Rel Pred"
    by blast
  from B1 have "S. (SourceTerm S, TargetTerm (S))  Rel¯"
    by simp
  moreover from B2 have "rel_preserves_binary_pred (Rel¯) Pred"
    by simp
  ultimately
  show "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_preserves_binary_pred Rel Pred"
    by blast
qed

lemma (in encoding) indRelR_cond_reflection_iff_indRelL_cond_preservation:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "(Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_reflects_pred Rel Pred)
         = (Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_preserves_pred Rel Pred)"
proof
  assume "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_reflects_pred Rel Pred"
  then obtain Rel where A1: "S. (SourceTerm S, TargetTerm (S))  Rel"
                    and A2: "rel_reflects_pred Rel Pred"
    by blast
  from A1 have "S. (TargetTerm (S), SourceTerm S)  Rel¯"
    by simp
  moreover from A2 have "rel_preserves_pred (Rel¯) Pred"
    by blast
  ultimately
  show "Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_preserves_pred Rel Pred"
    by blast
next
  assume "Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_preserves_pred Rel Pred"
  then obtain Rel where B1: "S. (TargetTerm (S), SourceTerm S)  Rel"
                    and B2: "rel_preserves_pred Rel Pred"
    by blast
  from B1 have "S. (SourceTerm S, TargetTerm (S))  Rel¯"
    by simp
  moreover from B2 have "rel_reflects_pred (Rel¯) Pred"
    by simp
  ultimately
  show "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_reflects_pred Rel Pred"
    by blast
qed

lemma (in encoding) indRelR_cond_binary_reflection_iff_indRelL_cond_binary_preservation:
  fixes Pred :: "('procS, 'procT) Proc  'b  bool"
  shows "(Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_reflects_binary_pred Rel Pred)
         = (Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)
             rel_preserves_binary_pred Rel Pred)"
proof
  assume "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_reflects_binary_pred Rel Pred"
  then obtain Rel where A1: "S. (SourceTerm S, TargetTerm (S))  Rel"
                    and A2: "rel_reflects_binary_pred Rel Pred"
    by blast
  from A1 have "S. (TargetTerm (S), SourceTerm S)  Rel¯"
    by simp
  moreover from A2 have "rel_preserves_binary_pred (Rel¯) Pred"
    by blast
  ultimately
  show "Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_preserves_binary_pred Rel Pred"
    by blast
next
  assume "Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_preserves_binary_pred Rel Pred"
  then obtain Rel where B1: "S. (TargetTerm (S), SourceTerm S)  Rel"
                    and B2: "rel_preserves_binary_pred Rel Pred"
    by blast
  from B1 have "S. (SourceTerm S, TargetTerm (S))  Rel¯"
    by simp
  moreover from B2 have "rel_reflects_binary_pred (Rel¯) Pred"
    by simp
  ultimately
  show "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_reflects_binary_pred Rel Pred"
    by blast
qed

lemma (in encoding) indRelR_cond_respection_iff_indRelL_cond_respection:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "(Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_respects_pred Rel Pred)
         = (Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_respects_pred Rel Pred)"
proof
  assume "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_respects_pred Rel Pred"
  from this obtain Rel where A1: "S. (SourceTerm S, TargetTerm (S))  Rel"
                         and A2: "rel_respects_pred Rel Pred"
    by blast
  from A1 have "S. (TargetTerm (S), SourceTerm S)  {(a, b). (b, a)  Rel}"
    by simp
  moreover from A2 have "rel_respects_pred {(a, b). (b, a)  Rel} Pred"
    by blast
  ultimately show "Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_respects_pred Rel Pred"
    by blast
next
  assume "Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_respects_pred Rel Pred"
  from this obtain Rel where A1: "S. (TargetTerm (S), SourceTerm S)  Rel"
                         and A2: "rel_respects_pred Rel Pred"
    by blast
  from A1 have "S. (SourceTerm S, TargetTerm (S))  {(a, b). (b, a)  Rel}"
    by simp
  moreover from A2 have "rel_respects_pred {(a, b). (b, a)  Rel} Pred"
    by blast
  ultimately show "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_respects_pred Rel Pred"
    by blast
qed

lemma (in encoding) indRelR_cond_binary_respection_iff_indRelL_cond_binary_respection:
  fixes Pred :: "('procS, 'procT) Proc  'b  bool"
  shows "(Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_respects_binary_pred Rel Pred)
         = (Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)
             rel_respects_binary_pred Rel Pred)"
proof
  assume "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_respects_binary_pred Rel Pred"
  from this obtain Rel where A1: "S. (SourceTerm S, TargetTerm (S))  Rel"
                         and A2: "rel_respects_binary_pred Rel Pred"
    by blast
  from A1 have "S. (TargetTerm (S), SourceTerm S)  {(a, b). (b, a)  Rel}"
    by simp
  moreover from A2 have "rel_respects_binary_pred {(a, b). (b, a)  Rel} Pred"
    by blast
  ultimately
  show "Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_respects_binary_pred Rel Pred"
    by blast
next
  assume "Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_respects_binary_pred Rel Pred"
  from this obtain Rel where A1: "S. (TargetTerm (S), SourceTerm S)  Rel"
                         and A2: "rel_respects_binary_pred Rel Pred"
    by blast
  from A1 have "S. (SourceTerm S, TargetTerm (S))  {(a, b). (b, a)  Rel}"
    by simp
  moreover from A2 have "rel_respects_binary_pred {(a, b). (b, a)  Rel} Pred"
    by blast
  ultimately
  show "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_respects_binary_pred Rel Pred"
    by blast
qed

text ‹An encoding preserves, reflects, or respects a predicate iff indRelL reflects, preserves,
        or respects this predicate.›

lemma (in encoding) enc_preserves_pred_iff_indRelL_reflects_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_preserves_pred Pred = rel_reflects_pred indRelL Pred"
      using enc_preserves_pred_iff_indRelR_preserves_pred[where Pred="Pred"]
            indRelR_preserves_pred_iff_indRelL_reflects_pred[where Pred="Pred"]
    by blast

lemma (in encoding) enc_reflects_pred_iff_indRelL_preserves_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_reflects_pred Pred = rel_preserves_pred indRelL Pred"
      using enc_reflects_pred_iff_indRelR_reflects_pred[where Pred="Pred"]
            indRelR_reflects_pred_iff_indRelL_preserves_pred[where Pred="Pred"]
    by blast

lemma (in encoding) enc_respects_pred_iff_indRelL_respects_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_respects_pred Pred = rel_respects_pred indRelL Pred"
      using enc_preserves_pred_iff_indRelL_reflects_pred[where Pred="Pred"]
            enc_reflects_pred_iff_indRelL_preserves_pred[where Pred="Pred"]
    by blast

text ‹An encoding preserves, reflects, or respects a predicate iff there exists a relation,
        namely indRelL, that relates literal translations with their source terms and reflects,
        preserves, or respects this predicate.›

lemma (in encoding) enc_preserves_pred_iff_source_target_rel_reflects_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_preserves_pred Pred
         = (Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_reflects_pred Rel Pred)"
      using enc_preserves_pred_iff_source_target_rel_preserves_pred[where Pred="Pred"]
            indRelR_cond_preservation_iff_indRelL_cond_reflection[where Pred="Pred"]
    by simp

lemma (in encoding) enc_reflects_pred_iff_source_target_rel_preserves_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_reflects_pred Pred
         = (Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_preserves_pred Rel Pred)"
      using enc_reflects_pred_iff_source_target_rel_reflects_pred[where Pred="Pred"]
            indRelR_cond_reflection_iff_indRelL_cond_preservation[where Pred="Pred"]
    by simp

lemma (in encoding) enc_respects_pred_iff_source_target_rel_respects_pred_encL:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_respects_pred Pred
         = (Rel. (S. (TargetTerm (S), SourceTerm S)  Rel)  rel_respects_pred Rel Pred)"
      using enc_respects_pred_iff_source_target_rel_respects_pred_encR[where Pred="Pred"]
            indRelR_cond_respection_iff_indRelL_cond_respection[where Pred="Pred"]
    by simp

text ‹To analyse the respection of source term behaviours we use relations that contain both kind
        of pairs: (S, enc S) as well as (enc S, S) for all source terms S.›

inductive_set (in encoding) indRel
    :: "((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
  where
  encR: "(SourceTerm S, TargetTerm (S))  indRel" |
  encL: "(TargetTerm (S), SourceTerm S)  indRel"

abbreviation (in encoding) indRelInfix ::
    "('procS, 'procT) Proc  ('procS, 'procT) Proc  bool" ("_ ℛ⟦⋅⟧ _" [75, 75] 80)
  where
  "P ℛ⟦⋅⟧ Q  (P, Q)  indRel"

lemma (in encoding) indRel_symm:
  shows "sym indRel"
      unfolding sym_def
    by (auto simp add: indRel.simps indRel.encR indRel.encL)

inductive_set (in encoding) indRelEQ
    :: "((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
  where
  encR:   "(SourceTerm S, TargetTerm (S))  indRelEQ" |
  encL:   "(TargetTerm (S), SourceTerm S)  indRelEQ" |
  target: "(TargetTerm T, TargetTerm T)  indRelEQ" |
  trans:  "(P, Q)  indRelEQ; (Q, R)  indRelEQ  (P, R)  indRelEQ"

abbreviation (in encoding) indRelEQinfix ::
    "('procS, 'procT) Proc  ('procS, 'procT) Proc  bool" ("_ ∼⟦⋅⟧ _" [75, 75] 80)
  where
  "P ∼⟦⋅⟧ Q  (P, Q)  indRelEQ"

lemma (in encoding) indRelEQ_refl:
  shows "refl indRelEQ"
    unfolding refl_on_def
proof auto
  fix P
  show "P ∼⟦⋅⟧ P"
  proof (cases P)
    case (SourceTerm SP)
    assume "SP ∈S P"
    moreover have "SourceTerm SP ∼⟦⋅⟧ TargetTerm (SP)"
      by (rule indRelEQ.encR)
    moreover have "TargetTerm (SP) ∼⟦⋅⟧ SourceTerm SP"
      by (rule indRelEQ.encL)
    ultimately show "P ∼⟦⋅⟧ P"
      by (simp add: indRelEQ.trans[where P="SourceTerm SP" and Q="TargetTerm (SP)"])
  next
    case (TargetTerm TP)
    assume "TP ∈T P"
    thus "P ∼⟦⋅⟧ P"
      by (simp add: indRelEQ.target)
  qed
qed

lemma (in encoding) indRelEQ_is_preorder:
  shows "preorder indRelEQ"
    unfolding preorder_on_def
proof
  show "refl indRelEQ"
    by (rule indRelEQ_refl)
next
  show "trans indRelEQ"
    unfolding trans_def
  proof clarify
    fix P Q R
    assume "P ∼⟦⋅⟧ Q" and "Q ∼⟦⋅⟧ R"
    thus "P ∼⟦⋅⟧ R"
      by (rule indRelEQ.trans)
  qed
qed

lemma (in encoding) indRelEQ_symm:
  shows "sym indRelEQ"
    unfolding sym_def
proof clarify
  fix P Q
  assume "P ∼⟦⋅⟧ Q"
  thus "Q ∼⟦⋅⟧ P"
  proof induct
    case (encR S)
    show "TargetTerm (S) ∼⟦⋅⟧ SourceTerm S"
      by (rule indRelEQ.encL)
  next
    case (encL S)
    show "SourceTerm S ∼⟦⋅⟧ TargetTerm (S)"
      by (rule indRelEQ.encR)
  next
    case (target T)
    show "TargetTerm T ∼⟦⋅⟧ TargetTerm T"
      by (rule indRelEQ.target)
  next
    case (trans P Q R)
    assume "R ∼⟦⋅⟧ Q" and "Q ∼⟦⋅⟧ P"
    thus "R ∼⟦⋅⟧ P"
      by (rule indRelEQ.trans)
  qed
qed

lemma (in encoding) indRelEQ_is_equivalence:
  shows "equivalence indRelEQ"
      using indRelEQ_is_preorder indRelEQ_symm
      unfolding equiv_def preorder_on_def
    by blast

lemma (in encoding) refl_trans_closure_of_indRel:
  shows "indRelEQ = indRel*"
proof auto
  fix P Q
  assume "P ∼⟦⋅⟧ Q"
  thus "(P, Q)  indRel*"
  proof induct
    case (encR S)
    show "(SourceTerm S, TargetTerm (S))  indRel*"
        using indRel.encR[of S]
      by simp
  next
    case (encL S)
    show "(TargetTerm (S), SourceTerm S)  indRel*"
        using indRel.encL[of S]
      by simp
  next
    case (target T)
    show "(TargetTerm T, TargetTerm T)  indRel*"
      by simp
  next
    case (trans P Q R)
    assume "(P, Q)  indRel*" and "(Q, R)  indRel*"
    thus "(P, R)  indRel*"
      by simp
  qed
next
  fix P Q
  assume "(P, Q)  indRel*"
  thus "P ∼⟦⋅⟧ Q"
  proof induct
    show "P ∼⟦⋅⟧ P"
        using indRelEQ_refl
        unfolding refl_on_def
      by simp
  next
    case (step Q R)
    assume "P ∼⟦⋅⟧ Q"
    moreover assume "Q ℛ⟦⋅⟧ R"
    hence "Q ∼⟦⋅⟧ R"
      by (induct, simp_all add: indRelEQ.encR indRelEQ.encL)
    ultimately show "P ∼⟦⋅⟧ R"
      by (rule indRelEQ.trans)
  qed
qed

lemma (in encoding) refl_symm_trans_closure_of_indRel:
  shows "indRelEQ = (symcl (indRel=))+"
proof -
  have "(symcl (indRel=))+ = (symcl indRel)*"
    by (rule refl_symm_trans_closure_is_symm_refl_trans_closure[where Rel="indRel"])
  moreover have "symcl indRel = indRel"
    by (simp add: indRel_symm symm_closure_of_symm_rel[where Rel="indRel"])
  ultimately show "indRelEQ = (symcl (indRel=))+"
    by (simp add: refl_trans_closure_of_indRel)
qed

lemma (in encoding) symm_closure_of_indRelR:
  shows "indRel = symcl indRelR"
    and "indRelEQ = (symcl (indRelR=))+"
proof -
  show "indRel = symcl indRelR"
  proof auto
    fix P Q
    assume "P ℛ⟦⋅⟧ Q"
    thus "(P, Q)  symcl indRelR"
      by (induct, simp_all add: symcl_def indRelR.encR)
  next
    fix P Q
    assume "(P, Q)  symcl indRelR"
    thus "P ℛ⟦⋅⟧ Q"
      by (auto simp add: symcl_def indRelR.simps indRel.encR indRel.encL)
  qed
  thus "indRelEQ = (symcl (indRelR=))+"
      using refl_symm_trans_closure_is_symm_refl_trans_closure[where Rel="indRelR"]
            refl_trans_closure_of_indRel
    by simp
qed

lemma (in encoding) symm_closure_of_indRelL:
  shows "indRel = symcl indRelL"
    and "indRelEQ = (symcl (indRelL=))+"
proof -
  show "indRel = symcl indRelL"
  proof auto
    fix P Q
    assume "P ℛ⟦⋅⟧ Q"
    thus "(P, Q)  symcl indRelL"
     by (induct, simp_all add: symcl_def indRelL.encL)
  next
    fix P Q
    assume "(P, Q)  symcl indRelL"
    thus "P ℛ⟦⋅⟧ Q"
      by (auto simp add: symcl_def indRelL.simps indRel.encR indRel.encL)
  qed
  thus "indRelEQ = (symcl (indRelL=))+"
      using refl_symm_trans_closure_is_symm_refl_trans_closure[where Rel="indRelL"]
            refl_trans_closure_of_indRel
    by simp
qed

text ‹The relation indRel is a combination of indRelL and indRelR. indRel respects a predicate
        iff indRelR (or indRelL) respects it.›

lemma (in encoding) indRel_respects_pred_iff_indRelR_respects_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "rel_respects_pred indRel Pred = rel_respects_pred indRelR Pred"
proof
  assume respection: "rel_respects_pred indRel Pred"
  show "rel_respects_pred indRelR Pred"
  proof auto
    fix P Q
    assume "P ℛ⟦⋅⟧R Q"
    from this obtain S where "S ∈S P" and "S ∈T Q"
      by (induct, blast)
    hence "P ℛ⟦⋅⟧ Q"
      by (simp add: indRel.encR)
    moreover assume "Pred P"
    ultimately show "Pred Q"
        using respection
      by blast
  next
    fix P Q
    assume "P ℛ⟦⋅⟧R Q"
    from this obtain S where "S ∈S P" and "S ∈T Q"
      by (induct, blast)
    hence "P ℛ⟦⋅⟧ Q"
      by (simp add: indRel.encR)
    moreover assume "Pred Q"
    ultimately show "Pred P"
        using respection
      by blast
  qed
next
  assume "rel_respects_pred indRelR Pred"
  thus "rel_respects_pred indRel Pred"
      using symm_closure_of_indRelR(1)
            respection_and_closures(2)[where Rel="indRelR" and Pred="Pred"]
    by blast
qed

lemma (in encoding) indRel_respects_binary_pred_iff_indRelR_respects_binary_pred:
  fixes Pred :: "('procS, 'procT) Proc  'b  bool"
  shows "rel_respects_binary_pred indRel Pred = rel_respects_binary_pred indRelR Pred"
proof
  assume respection: "rel_respects_binary_pred indRel Pred"
  show "rel_respects_binary_pred indRelR Pred"
  proof auto
    fix P Q x
    assume "P ℛ⟦⋅⟧R Q"
    from this obtain S where "S ∈S P" and "S ∈T Q"
      by (induct, blast)
    hence "P ℛ⟦⋅⟧ Q"
      by (simp add: indRel.encR)
    moreover assume "Pred P x"
    ultimately show "Pred Q x"
        using respection
      by blast
  next
    fix P Q x
    assume "P ℛ⟦⋅⟧R Q"
    from this obtain S where "S ∈S P" and "S ∈T Q"
      by (induct, blast)
    hence "P ℛ⟦⋅⟧ Q"
      by (simp add: indRel.encR)
    moreover assume "Pred Q x"
    ultimately show "Pred P x"
        using respection
      by blast
  qed
next
  assume "rel_respects_binary_pred indRelR Pred"
  thus "rel_respects_binary_pred indRel Pred"
      using symm_closure_of_indRelR(1)
            respection_of_binary_predicates_and_closures(2)[where Rel="indRelR" and Pred="Pred"]
    by blast
qed

lemma (in encoding) indRel_cond_respection_iff_indRelR_cond_respection:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "(Rel.
          (S. (SourceTerm S, TargetTerm (S))  Rel  (TargetTerm (S), SourceTerm S)  Rel)
           rel_respects_pred Rel Pred)
         = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_respects_pred Rel Pred)"
proof
  assume "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel
           (TargetTerm (S), SourceTerm S)  Rel)  rel_respects_pred Rel Pred"
  from this obtain Rel
    where "S. (SourceTerm S, TargetTerm (S))  Rel  (TargetTerm (S), SourceTerm S)  Rel"
    and "rel_respects_pred Rel Pred"
    by blast
  thus "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_respects_pred Rel Pred"
    by blast
next
  assume "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_respects_pred Rel Pred"
  from this obtain Rel where A1: "S. (SourceTerm S, TargetTerm (S))  Rel"
                         and A2: "rel_respects_pred Rel Pred"
    by blast
  from A1 have "S. (SourceTerm S, TargetTerm (S))  symcl Rel
                 (TargetTerm (S), SourceTerm S)  symcl Rel"
    by (simp add: symcl_def)
  moreover from A2 have "rel_respects_pred (symcl Rel) Pred"
      using respection_and_closures(2)[where Rel="Rel" and Pred="Pred"]
    by blast
  ultimately
  show "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel  (TargetTerm (S), SourceTerm S)  Rel)
         rel_respects_pred Rel Pred"
    by blast
qed

lemma (in encoding) indRel_cond_binary_respection_iff_indRelR_cond_binary_respection:
  fixes Pred :: "('procS, 'procT) Proc  'b  bool"
  shows "(Rel.
          (S. (SourceTerm S, TargetTerm (S))  Rel  (TargetTerm (S), SourceTerm S)  Rel)
           rel_respects_binary_pred Rel Pred)
         = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)
             rel_respects_binary_pred Rel Pred)"
proof
  assume "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel
           (TargetTerm (S), SourceTerm S)  Rel)  rel_respects_binary_pred Rel Pred"
  from this obtain Rel
    where "S. (SourceTerm S, TargetTerm (S))  Rel  (TargetTerm (S), SourceTerm S)  Rel"
    and "rel_respects_binary_pred Rel Pred"
    by blast
  thus "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_respects_binary_pred Rel Pred"
    by blast
next
  assume "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  rel_respects_binary_pred Rel Pred"
  from this obtain Rel where A1: "S. (SourceTerm S, TargetTerm (S))  Rel"
                         and A2: "rel_respects_binary_pred Rel Pred"
    by blast
  from A1 have "S. (SourceTerm S, TargetTerm (S))  symcl Rel
                 (TargetTerm (S), SourceTerm S)  symcl Rel"
    by (simp add: symcl_def)
  moreover from A2 have "rel_respects_binary_pred (symcl Rel) Pred"
      using respection_of_binary_predicates_and_closures(2)[where Rel="Rel" and Pred="Pred"]
    by blast
  ultimately
  show "Rel. (S. (SourceTerm S, TargetTerm (S))  Rel  (TargetTerm (S), SourceTerm S)  Rel)
         rel_respects_binary_pred Rel Pred"
    by blast
qed

text ‹An encoding respects a predicate iff indRel respects this predicate.›

lemma (in encoding) enc_respects_pred_iff_indRel_respects_pred:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_respects_pred Pred = rel_respects_pred indRel Pred"
      using enc_respects_pred_iff_indRelR_respects_pred[where Pred="Pred"]
            indRel_respects_pred_iff_indRelR_respects_pred[where Pred="Pred"]
    by simp

text ‹An encoding respects a predicate iff there exists a relation, namely indRel, that relates
        source terms and their literal translations in both directions and respects this predicate.
›

lemma (in encoding) enc_respects_pred_iff_source_target_rel_respects_pred_encRL:
  fixes Pred :: "('procS, 'procT) Proc  bool"
  shows "enc_respects_pred Pred
         = (Rel.
            (S. (SourceTerm S, TargetTerm (S))  Rel  (TargetTerm (S), SourceTerm S)  Rel)
             rel_respects_pred Rel Pred)"
      using enc_respects_pred_iff_source_target_rel_respects_pred_encR[where Pred="Pred"]
            indRel_cond_respection_iff_indRelR_cond_respection[where Pred="Pred"]
    by simp

subsection ‹Relations Induced by the Encoding and a Relation on Target Terms›

text ‹Some encodability like e.g. operational correspondence are defined w.r.t. a relation on
        target terms. To analyse such criteria we include the respective target term relation in
        the considered relation on the disjoint union of source and target terms.›

inductive_set (in encoding) indRelRT
    :: "('procT × 'procT) set  ((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
    for TRel :: "('procT × 'procT) set"
  where
  encR:   "(SourceTerm S, TargetTerm (S))  indRelRT TRel" |
  target: "(T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  indRelRT TRel"

abbreviation (in encoding) indRelRTinfix
    :: "('procS, 'procT) Proc  ('procT × 'procT) set  ('procS, 'procT) Proc  bool"
       ("_ ℛ⟦⋅⟧RT<_> _" [75, 75, 75] 80)
  where
  "P ℛ⟦⋅⟧RT<TRel> Q  (P, Q)  indRelRT TRel"

inductive_set (in encoding) indRelRTPO
    :: "('procT × 'procT) set  ((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
    for TRel :: "('procT × 'procT) set"
  where
  encR:   "(SourceTerm S, TargetTerm (S))  indRelRTPO TRel" |
  source: "(SourceTerm S, SourceTerm S)  indRelRTPO TRel" |
  target: "(T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  indRelRTPO TRel" |
  trans:  "(P, Q)  indRelRTPO TRel; (Q, R)  indRelRTPO TRel  (P, R)  indRelRTPO TRel"

abbreviation (in encoding) indRelRTPOinfix
    :: "('procS, 'procT) Proc  ('procT × 'procT) set  ('procS, 'procT) Proc  bool"
       ("_ ≲⟦⋅⟧RT<_> _" [75, 75, 75] 80)
  where
  "P ≲⟦⋅⟧RT<TRel> Q  (P, Q)  indRelRTPO TRel"

lemma (in encoding) indRelRTPO_refl:
  fixes TRel :: "('procT × 'procT) set"
  assumes refl: "refl TRel"
  shows "refl (indRelRTPO TRel)"
    unfolding refl_on_def
proof auto
  fix P
  show "P ≲⟦⋅⟧RT<TRel> P"
  proof (cases P)
    case (SourceTerm SP)
    assume "SP ∈S P"
    thus "P ≲⟦⋅⟧RT<TRel> P"
      by (simp add: indRelRTPO.source)
  next
    case (TargetTerm TP)
    assume "TP ∈T P"
    with refl show "P ≲⟦⋅⟧RT<TRel> P"
        unfolding refl_on_def
      by (simp add: indRelRTPO.target)
  qed
qed

lemma (in encoding) refl_trans_closure_of_indRelRT:
  fixes TRel :: "('procT × 'procT) set"
  assumes refl: "refl TRel"
  shows "indRelRTPO TRel = (indRelRT TRel)*"
proof auto
  fix P Q
  assume "P ≲⟦⋅⟧RT<TRel> Q"
  thus "(P, Q)  (indRelRT TRel)*"
  proof induct
    case (encR S)
    show "(SourceTerm S, TargetTerm (S))  (indRelRT TRel)*"
        using indRelRT.encR[of S TRel]
      by simp
  next
    case (source S)
    show "(SourceTerm S, SourceTerm S)  (indRelRT TRel)*"
      by simp
  next
    case (target T1 T2)
    assume "(T1, T2)  TRel"
    thus "(TargetTerm T1, TargetTerm T2)  (indRelRT TRel)*"
        using indRelRT.target[of T1 T2 TRel]
      by simp
  next
    case (trans P Q R)
    assume "(P, Q)  (indRelRT TRel)*" and "(Q, R)  (indRelRT TRel)*"
    thus "(P, R)  (indRelRT TRel)*"
      by simp
  qed
next
  fix P Q
  assume "(P, Q)  (indRelRT TRel)*"
  thus "P ≲⟦⋅⟧RT<TRel> Q"
  proof induct
    from refl show "P ≲⟦⋅⟧RT<TRel> P"
        using indRelRTPO_refl[of TRel]
        unfolding refl_on_def
      by simp
  next
    case (step Q R)
    assume "P ≲⟦⋅⟧RT<TRel> Q"
    moreover assume "Q ℛ⟦⋅⟧RT<TRel> R"
    hence "Q ≲⟦⋅⟧RT<TRel> R"
      by (induct, simp_all add: indRelRTPO.encR indRelRTPO.target)
    ultimately show "P ≲⟦⋅⟧RT<TRel> R"
      by (rule indRelRTPO.trans)
  qed
qed

lemma (in encoding) indRelRTPO_is_preorder:
  fixes TRel :: "('procT × 'procT) set"
  assumes reflT: "refl TRel"
  shows "preorder (indRelRTPO TRel)"
    unfolding preorder_on_def
proof
  from reflT show "refl (indRelRTPO TRel)"
    by (rule indRelRTPO_refl)
next
  show "trans (indRelRTPO TRel)"
    unfolding trans_def
  proof clarify
    fix P Q R
    assume "P ≲⟦⋅⟧RT<TRel> Q" and "Q ≲⟦⋅⟧RT<TRel> R"
    thus "P ≲⟦⋅⟧RT<TRel> R"
        using indRelRTPO.trans
      by blast
  qed
qed

lemma (in encoding) transitive_closure_of_TRel_to_indRelRTPO:
  fixes TRel  :: "('procT × 'procT) set"
    and TP TQ :: "'procT"
  shows "(TP, TQ)  TRel+  TargetTerm TP ≲⟦⋅⟧RT<TRel> TargetTerm TQ"
proof -
  assume "(TP, TQ)  TRel+"
  thus "TargetTerm TP ≲⟦⋅⟧RT<TRel> TargetTerm TQ"
  proof induct
    fix TQ
    assume "(TP, TQ)  TRel"
    thus "TargetTerm TP ≲⟦⋅⟧RT<TRel> TargetTerm TQ"
      by (rule indRelRTPO.target)
  next
    case (step TQ TR)
    assume "TargetTerm TP ≲⟦⋅⟧RT<TRel> TargetTerm TQ"
    moreover assume "(TQ, TR)  TRel"
    hence "TargetTerm TQ ≲⟦⋅⟧RT<TRel> TargetTerm TR"
      by (simp add: indRelRTPO.target)
    ultimately show "TargetTerm TP ≲⟦⋅⟧RT<TRel> TargetTerm TR"
      by (rule indRelRTPO.trans)
  qed
qed

text ‹The relation indRelRT is the smallest relation that relates all source terms and their
        literal translations and contains TRel. Thus there exists a relation that relates source
        terms and their literal translations and satisfies some predicate on its pairs iff the
        predicate holds for the pairs of indRelR.›

lemma (in encoding) indRelR_modulo_pred_impl_indRelRT_modulo_pred:
  fixes Pred :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc)  bool"
  shows "((P, Q)  indRelR. Pred (P, Q)) = (TRel. ((TP, TQ)  TRel.
         Pred (TargetTerm TP, TargetTerm TQ))  ((P, Q)  indRelRT TRel. Pred (P, Q)))"
proof (rule iffI)
  assume A: "(P, Q)  indRelR. Pred (P, Q)"
  show "TRel. ((TP, TQ)  TRel. Pred (TargetTerm TP, TargetTerm TQ))
        = ((P, Q)  indRelRT TRel. Pred (P, Q))"
  proof (rule allI, rule iffI)
    fix TRel
    assume "(TP, TQ)  TRel. Pred (TargetTerm TP, TargetTerm TQ)"
    with A show "(P, Q)  indRelRT TRel. Pred (P, Q)"
      by (auto simp add: indRelR.encR indRelRT.simps)
  next
    fix TRel
    assume "(P, Q)  indRelRT TRel. Pred (P, Q)"
    thus "(TP, TQ)  TRel. Pred (TargetTerm TP, TargetTerm TQ)"
      by (auto simp add: indRelRT.target)
  qed
next
  assume "TRel. ((TP, TQ)  TRel. Pred (TargetTerm TP, TargetTerm TQ))
           ((P, Q)  indRelRT TRel. Pred (P, Q))"
  hence B: "TRel. ((TP, TQ)  TRel. Pred (TargetTerm TP, TargetTerm TQ))
             ((P, Q)  indRelRT TRel. Pred (P, Q))"
    by blast
  have "S. Pred (SourceTerm S, TargetTerm (S))"
      using B[of "{}"]
    by (simp add: indRelRT.simps)
  thus "(P, Q)  indRelR. Pred (P, Q)"
    by (auto simp add: indRelR.simps)
qed

lemma (in encoding) indRelRT_iff_exists_source_target_relation:
  fixes Pred :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc)  bool"
  shows "(TRel. ((TP, TQ)  TRel. Pred (TargetTerm TP, TargetTerm TQ))
           ((P, Q)  indRelRT TRel. Pred (P, Q)))
         = (Rel. (S. (SourceTerm S, TargetTerm (S))  Rel)  ((P, Q)  Rel. Pred (P, Q)))"
      using indRelR_iff_exists_source_target_relation[where Pred="Pred"]
            indRelR_modulo_pred_impl_indRelRT_modulo_pred[where Pred="Pred"]
    by simp

lemma (in encoding) indRelRT_modulo_pred_impl_indRelRTPO_modulo_pred:
  fixes TRel :: "('procT × 'procT) set"
    and Pred :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc)  bool"
  assumes reflCond:  "P. Pred (P, P)"
      and transCond: "P Q R. Pred (P, Q)  Pred (Q, R)  Pred (P, R)"
  shows "((P, Q)  indRelRT TRel. Pred (P, Q)) = ((P, Q)  indRelRTPO TRel. Pred (P, Q))"
proof auto
  fix P Q
  assume A: "x  indRelRT TRel. Pred x"
  assume "P ≲⟦⋅⟧RT<TRel> Q"
  thus "Pred (P, Q)"
  proof induct
    case (encR S)
    have "SourceTerm S ℛ⟦⋅⟧RT<TRel> TargetTerm (S)"
      by (simp add: indRelRT.encR)
    with A show "Pred (SourceTerm S, TargetTerm (S))"
      by simp
  next
    case (source S)
    from reflCond show "Pred (SourceTerm S, SourceTerm S)"
      by simp
  next
    case (target T1 T2)
    assume "(T1, T2)  TRel"
    hence "TargetTerm T1 ℛ⟦⋅⟧RT<TRel> TargetTerm T2"
      by (simp add: indRelRT.target)
    with A show "Pred (TargetTerm T1, TargetTerm T2)"
      by simp
  next
    case (trans P Q R)
    assume "Pred (P, Q)" and "Pred (Q, R)"
    with transCond show "Pred (P, R)"
      by blast
  qed
next
  fix P Q
  assume "x  indRelRTPO TRel. Pred x" and "P ℛ⟦⋅⟧RT<TRel> Q"
  thus "Pred (P, Q)"
    by (auto simp add: indRelRTPO.encR indRelRTPO.target indRelRT.simps)
qed

lemma (in encoding) indRelR_modulo_pred_impl_indRelRTPO_modulo_pred:
  fixes Pred :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc)  bool"
  assumes "P. Pred (P, P)"
      and "P Q R. Pred (P, Q)  Pred (Q, R)  Pred (P, R)"
  shows "((P, Q)  indRelR. Pred (P, Q))
         = (TRel. ((TP, TQ)  TRel. Pred (TargetTerm TP, TargetTerm TQ))
             ((P, Q)  indRelRTPO TRel. Pred (P, Q)))"
proof -
  have "((P, Q)indRelR. Pred (P, Q)) = (TRel. ((TP, TQ)  TRel.
        Pred (TargetTerm TP, TargetTerm TQ))  ((P, Q)  indRelRT TRel. Pred (P, Q)))"
      using indRelR_modulo_pred_impl_indRelRT_modulo_pred[where Pred="Pred"]
    by simp
  moreover
  have "TRel. ((P, Q)indRelRT TRel. Pred (P, Q)) = ((P, Q)indRelRTPO TRel. Pred (P, Q))"
      using assms indRelRT_modulo_pred_impl_indRelRTPO_modulo_pred[where Pred="Pred"]
    by blast
  ultimately show ?thesis
    by simp
qed

text ‹The relation indRelLT includes TRel and relates literal translations and their source
        terms.›

inductive_set (in encoding) indRelLT
    :: "('procT × 'procT) set  ((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
    for TRel :: "('procT × 'procT) set"
  where
  encL:   "(TargetTerm (S), SourceTerm S)  indRelLT TRel" |
  target: "(T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  indRelLT TRel"

abbreviation (in encoding) indRelLTinfix
    :: "('procS, 'procT) Proc  ('procT × 'procT) set  ('procS, 'procT) Proc  bool"
       ("_ ℛ⟦⋅⟧LT<_> _" [75, 75, 75] 80)
  where
  "P ℛ⟦⋅⟧LT<TRel> Q  (P, Q)  indRelLT TRel"

inductive_set (in encoding) indRelLTPO
    :: "('procT × 'procT) set  ((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
    for TRel :: "('procT × 'procT) set"
  where
  encL:   "(TargetTerm (S), SourceTerm S)  indRelLTPO TRel" |
  source: "(SourceTerm S, SourceTerm S)  indRelLTPO TRel" |
  target: "(T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  indRelLTPO TRel" |
  trans:  "(P, Q)  indRelLTPO TRel; (Q, R)  indRelLTPO TRel  (P, R)  indRelLTPO TRel"

abbreviation (in encoding) indRelLTPOinfix
    :: "('procS, 'procT) Proc  ('procT × 'procT) set  ('procS, 'procT) Proc  bool"
       ("_ ≲⟦⋅⟧LT<_> _" [75, 75, 75] 80)
  where
  "P ≲⟦⋅⟧LT<TRel> Q  (P, Q)  indRelLTPO TRel"

lemma (in encoding) indRelLTPO_refl:
  fixes TRel :: "('procT × 'procT) set"
  assumes refl: "refl TRel"
  shows "refl (indRelLTPO TRel)"
    unfolding refl_on_def
proof auto
  fix P
  show "P ≲⟦⋅⟧LT<TRel> P"
  proof (cases P)
    case (SourceTerm SP)
    assume "SP ∈S P"
    thus "P ≲⟦⋅⟧LT<TRel> P"
      by (simp add: indRelLTPO.source)
  next
    case (TargetTerm TP)
    assume "TP ∈T P"
    with refl show "P ≲⟦⋅⟧LT<TRel> P"
        using indRelLTPO.target[of TP TP TRel]
        unfolding refl_on_def
      by simp
  qed
qed

lemma (in encoding) refl_trans_closure_of_indRelLT:
  fixes TRel :: "('procT × 'procT) set"
  assumes refl: "refl TRel"
  shows "indRelLTPO TRel = (indRelLT TRel)*"
proof auto
  fix P Q
  assume "P ≲⟦⋅⟧LT<TRel> Q"
  thus "(P, Q)  (indRelLT TRel)*"
  proof induct
    case (encL S)
    show "(TargetTerm (S), SourceTerm S)  (indRelLT TRel)*"
        using indRelLT.encL[of S TRel]
      by simp
  next
    case (source S)
    show "(SourceTerm S, SourceTerm S)  (indRelLT TRel)*"
      by simp
  next
    case (target T1 T2)
    assume "(T1, T2)  TRel"
    thus "(TargetTerm T1, TargetTerm T2)  (indRelLT TRel)*"
        using indRelLT.target[of T1 T2 TRel]
      by simp
  next
    case (trans P Q R)
    assume "(P, Q)  (indRelLT TRel)*" and "(Q, R)  (indRelLT TRel)*"
    thus "(P, R)  (indRelLT TRel)*"
      by simp
  qed
next
  fix P Q
  assume "(P, Q)  (indRelLT TRel)*"
  thus "P ≲⟦⋅⟧LT<TRel> Q"
  proof induct
    from refl show "P ≲⟦⋅⟧LT<TRel> P"
        using indRelLTPO_refl[of TRel]
        unfolding refl_on_def
      by simp
  next
    case (step Q R)
    assume "P ≲⟦⋅⟧LT<TRel> Q"
    moreover assume "Q ℛ⟦⋅⟧LT<TRel> R"
    hence "Q ≲⟦⋅⟧LT<TRel> R"
      by (induct, simp_all add: indRelLTPO.encL indRelLTPO.target)
    ultimately show "P ≲⟦⋅⟧LT<TRel> R"
      by (rule indRelLTPO.trans)
  qed
qed

inductive_set (in encoding) indRelT
    :: "('procT × 'procT) set  ((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
    for TRel :: "('procT × 'procT) set"
  where
  encR:   "(SourceTerm S, TargetTerm (S))  indRelT TRel" |
  encL:   "(TargetTerm (S), SourceTerm S)  indRelT TRel" |
  target: "(T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  indRelT TRel"

abbreviation (in encoding) indRelTinfix
    :: "('procS, 'procT) Proc  ('procT × 'procT) set  ('procS, 'procT) Proc  bool"
       ("_ ℛ⟦⋅⟧T<_> _" [75, 75, 75] 80)
  where
  "P ℛ⟦⋅⟧T<TRel> Q  (P, Q)  indRelT TRel"

lemma (in encoding) indRelT_symm:
  fixes TRel :: "('procT × 'procT) set"
  assumes symm: "sym TRel"
  shows "sym (indRelT TRel)"
    unfolding sym_def
proof clarify
  fix P Q
  assume "(P, Q)  indRelT TRel"
  thus "(Q, P)  indRelT TRel"
      using symm
      unfolding sym_def
    by (induct, simp_all add: indRelT.encL indRelT.encR indRelT.target)
qed

inductive_set (in encoding) indRelTEQ
    :: "('procT × 'procT) set  ((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
    for TRel :: "('procT × 'procT) set"
  where
  encR:   "(SourceTerm S, TargetTerm (S))  indRelTEQ TRel" |
  encL:   "(TargetTerm (S), SourceTerm S)  indRelTEQ TRel" |
  target: "(T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  indRelTEQ TRel" |
  trans:  "(P, Q)  indRelTEQ TRel; (Q, R)  indRelTEQ TRel  (P, R)  indRelTEQ TRel"

abbreviation (in encoding) indRelTEQinfix
    :: "('procS, 'procT) Proc  ('procT × 'procT) set  ('procS, 'procT) Proc  bool"
       ("_ ∼⟦⋅⟧T<_> _" [75, 75, 75] 80)
  where
  "P ∼⟦⋅⟧T<TRel> Q  (P, Q)  indRelTEQ TRel"

lemma (in encoding) indRelTEQ_refl:
  fixes TRel :: "('procT × 'procT) set"
  assumes refl: "refl TRel"
  shows "refl (indRelTEQ TRel)"
    unfolding refl_on_def
proof auto
  fix P
  show "P ∼⟦⋅⟧T<TRel> P"
  proof (cases P)
    case (SourceTerm SP)
    assume "SP ∈S P"
    moreover have "SourceTerm SP ∼⟦⋅⟧T<TRel> TargetTerm (SP)"
      by (rule indRelTEQ.encR)
    moreover have "TargetTerm (SP) ∼⟦⋅⟧T<TRel> SourceTerm SP"
      by (rule indRelTEQ.encL)
    ultimately show "P ∼⟦⋅⟧T<TRel> P"
      by (simp add: indRelTEQ.trans[where P="SourceTerm SP" and Q="TargetTerm (SP)"])
  next
    case (TargetTerm TP)
    assume "TP ∈T P"
    with refl show "P ∼⟦⋅⟧T<TRel> P"
        unfolding refl_on_def
      by (simp add: indRelTEQ.target)
  qed
qed

lemma (in encoding) indRelTEQ_symm:
  fixes TRel :: "('procT × 'procT) set"
  assumes symm: "sym TRel"
  shows "sym (indRelTEQ TRel)"
    unfolding sym_def
proof clarify
  fix P Q
  assume "P ∼⟦⋅⟧T<TRel> Q"
  thus "Q ∼⟦⋅⟧T<TRel> P"
  proof induct
    case (encR S)
    show "TargetTerm (S) ∼⟦⋅⟧T<TRel> SourceTerm S"
      by (rule indRelTEQ.encL)
  next
    case (encL S)
    show "SourceTerm S ∼⟦⋅⟧T<TRel> TargetTerm (S)"
      by (rule indRelTEQ.encR)
  next
    case (target T1 T2)
    assume "(T1, T2)  TRel"
    with symm show "TargetTerm T2 ∼⟦⋅⟧T<TRel> TargetTerm T1"
        unfolding sym_def
      by (simp add: indRelTEQ.target)
  next
    case (trans P Q R)
    assume "R ∼⟦⋅⟧T<TRel> Q" and "Q ∼⟦⋅⟧T<TRel> P"
    thus "R ∼⟦⋅⟧T<TRel> P"
      by (rule indRelTEQ.trans)
  qed
qed

lemma (in encoding) refl_trans_closure_of_indRelT:
  fixes TRel :: "('procT × 'procT) set"
  assumes refl: "refl TRel"
  shows "indRelTEQ TRel = (indRelT TRel)*"
proof auto
  fix P Q
  assume "P ∼⟦⋅⟧T<TRel> Q"
  thus "(P, Q)  (indRelT TRel)*"
  proof induct
    case (encR S)
    show "(SourceTerm S, TargetTerm (S))  (indRelT TRel)*"
        using indRelT.encR[of S TRel]
      by simp
  next
    case (encL S)
    show "(TargetTerm (S), SourceTerm S)  (indRelT TRel)*"
        using indRelT.encL[of S TRel]
      by simp
  next
    case (target T1 T2)
    assume "(T1, T2)  TRel"
    thus "(TargetTerm T1, TargetTerm T2)  (indRelT TRel)*"
        using indRelT.target[of T1 T2 TRel]
      by simp
  next
    case (trans P Q R)
    assume "(P, Q)  (indRelT TRel)*" and "(Q, R)  (indRelT TRel)*"
    thus "(P, R)  (indRelT TRel)*"
      by simp
  qed
next
  fix P Q
  assume "(P, Q)  (indRelT TRel)*"
  thus "P ∼⟦⋅⟧T<TRel> Q"
  proof induct
    from refl show "P ∼⟦⋅⟧T<TRel> P"
        using indRelTEQ_refl[of TRel]
        unfolding refl_on_def
      by simp
  next
    case (step Q R)
    assume "P ∼⟦⋅⟧T<TRel> Q"
    moreover assume "Q ℛ⟦⋅⟧T<TRel> R"
    hence "Q ∼⟦⋅⟧T<TRel> R"
      by (induct, simp_all add: indRelTEQ.encR indRelTEQ.encL indRelTEQ.target)
    ultimately show "P ∼⟦⋅⟧T<TRel> R"
      by (rule indRelTEQ.trans)
  qed
qed

lemma (in encoding) refl_symm_trans_closure_of_indRelT:
  fixes TRel :: "('procT × 'procT) set"
  assumes refl: "refl TRel"
      and symm: "sym TRel"
  shows "indRelTEQ TRel = (symcl ((indRelT TRel)=))+"
proof -
  have "(symcl ((indRelT TRel)=))+ = (symcl (indRelT TRel))*"
    by (rule refl_symm_trans_closure_is_symm_refl_trans_closure[where Rel="indRelT TRel"])
  moreover from symm have "symcl (indRelT TRel) = indRelT TRel"
      using indRelT_symm[where TRel="TRel"] symm_closure_of_symm_rel[where Rel="indRelT TRel"]
    by blast
  ultimately show "indRelTEQ TRel = (symcl ((indRelT TRel)=))+"
      using refl refl_trans_closure_of_indRelT[where TRel="TRel"]
    by simp
qed

lemma (in encoding) symm_closure_of_indRelRT:
  fixes TRel :: "('procT × 'procT) set"
  assumes refl: "refl TRel"
      and symm: "sym TRel"
  shows "indRelT TRel = symcl (indRelRT TRel)"
    and "indRelTEQ TRel = (symcl ((indRelRT TRel)=))+"
proof -
  show "indRelT TRel = symcl (indRelRT TRel)"
  proof auto
    fix P Q
    assume "P ℛ⟦⋅⟧T<TRel> Q"
    thus "(P, Q)  symcl (indRelRT TRel)"
      by (induct, simp_all add: symcl_def indRelRT.encR indRelRT.target)
  next
    fix P Q
    assume "(P, Q)  symcl (indRelRT TRel)"
    thus "P ℛ⟦⋅⟧T<TRel> Q"
    proof (auto simp add: symcl_def indRelRT.simps)
      fix S
      show "SourceTerm S ℛ⟦⋅⟧T<TRel> TargetTerm (S)"
        by (rule indRelT.encR)
    next
      fix T1 T2
      assume "(T1, T2)  TRel"
      thus "TargetTerm T1 ℛ⟦⋅⟧T<TRel> TargetTerm T2"
        by (rule indRelT.target)
    next
      fix S
      show "TargetTerm (S) ℛ⟦⋅⟧T<TRel> SourceTerm S"
        by (rule indRelT.encL)
    next
      fix T1 T2
      assume "(T1, T2)  TRel"
      with symm show "TargetTerm T2 ℛ⟦⋅⟧T<TRel> TargetTerm T1"
          unfolding sym_def
        by (simp add: indRelT.target)
    qed
  qed
  with refl show "indRelTEQ TRel = (symcl ((indRelRT TRel)=))+"
      using refl_symm_trans_closure_is_symm_refl_trans_closure[where Rel="indRelRT TRel"]
            refl_trans_closure_of_indRelT
    by simp
qed

lemma (in encoding) symm_closure_of_indRelLT:
  fixes TRel :: "('procT × 'procT) set"
  assumes refl: "refl TRel"
      and symm: "sym TRel"
  shows "indRelT TRel = symcl (indRelLT TRel)"
    and "indRelTEQ TRel = (symcl ((indRelLT TRel)=))+"
proof -
  show "indRelT TRel = symcl (indRelLT TRel)"
  proof auto
    fix P Q
    assume "P ℛ⟦⋅⟧T<TRel> Q"
    thus "(P, Q)  symcl (indRelLT TRel)"
      by (induct, simp_all add: symcl_def indRelLT.encL indRelLT.target)
  next
    fix P Q
    assume "(P, Q)  symcl (indRelLT TRel)"
    thus "P ℛ⟦⋅⟧T<TRel> Q"
    proof (auto simp add: symcl_def indRelLT.simps)
      fix S
      show "SourceTerm S ℛ⟦⋅⟧T<TRel> TargetTerm (S)"
        by (rule indRelT.encR)
    next
      fix T1 T2
      assume "(T1, T2)  TRel"
      thus "TargetTerm T1 ℛ⟦⋅⟧T<TRel> TargetTerm T2"
        by (rule indRelT.target)
    next
      fix S
      show "TargetTerm (S) ℛ⟦⋅⟧T<TRel> SourceTerm S"
        by (rule indRelT.encL)
    next
      fix T1 T2
      assume "(T1, T2)  TRel"
      with symm show "TargetTerm T2 ℛ⟦⋅⟧T<TRel> TargetTerm T1"
          unfolding sym_def
        by (simp add: indRelT.target)
    qed
  qed
  with refl show "indRelTEQ TRel = (symcl ((indRelLT TRel)=))+"
      using refl_symm_trans_closure_is_symm_refl_trans_closure[where Rel="indRelLT TRel"]
            refl_trans_closure_of_indRelT
    by simp
qed

text ‹If the relations indRelRT, indRelLT, or indRelT contain a pair of target terms, then this
        pair is also related by the considered target term relation.›

lemma (in encoding) indRelRT_to_TRel:
  fixes TRel  :: "('procT × 'procT) set"
    and TP TQ :: "'procT"
  assumes rel: "TargetTerm TP ℛ⟦⋅⟧RT<TRel> TargetTerm TQ"
  shows "(TP, TQ)  TRel"
      using rel
    by (simp add: indRelRT.simps)

lemma (in encoding) indRelLT_to_TRel:
  fixes TRel  :: "('procT × 'procT) set"
    and TP TQ :: "'procT"
  assumes rel: "TargetTerm TP ℛ⟦⋅⟧LT<TRel> TargetTerm TQ"
  shows "(TP, TQ)  TRel"
      using rel
    by (simp add: indRelLT.simps)

lemma (in encoding) indRelT_to_TRel:
  fixes TRel  :: "('procT × 'procT) set"
    and TP TQ :: "'procT"
  assumes rel: "TargetTerm TP ℛ⟦⋅⟧T<TRel> TargetTerm TQ"
  shows "(TP, TQ)  TRel"
      using rel
    by (simp add: indRelT.simps)

text ‹If the preorders indRelRTPO, indRelLTPO, or the equivalence indRelTEQ contain a pair of
        terms, then the pair of target terms that is related to these two terms is also related by
        the reflexive and transitive closure of the considered target term relation.›

lemma (in encoding) indRelRTPO_to_TRel:
  fixes TRel :: "('procT × 'procT) set"
    and P Q  :: "('procS, 'procT) Proc"
  assumes rel: "P ≲⟦⋅⟧RT<TRel> Q"
  shows "SP SQ. SP ∈S P  SQ ∈S Q  SP = SQ"
    and "SP TQ. SP ∈S P  TQ ∈T Q
          (SP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
    and "TP SQ. TP ∈T P  SQ ∈S Q  False"
    and "TP TQ. TP ∈T P  TQ ∈T Q  (TP, TQ)  TRel+"
proof -
  have reflTRel: "S. (S, S)  TRel  {(T1, T2). S. T1 = S  T2 = S}"
    by auto
  from rel show "SP SQ. SP ∈S P  SQ ∈S Q  SP = SQ"
            and "SP TQ. SP ∈S P  TQ ∈T Q
                  (SP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
            and "TP SQ. TP ∈T P  SQ ∈S Q  False"
            and "TP TQ. TP ∈T P  TQ ∈T Q  (TP, TQ)  TRel+"
  proof induct
    case (encR S)
    show "SP SQ. SP ∈S SourceTerm S  SQ ∈S TargetTerm (S)  SP = SQ"
     and "TP SQ. TP ∈T SourceTerm S  SQ ∈S TargetTerm (S)  False"
     and "TP TQ. TP ∈T SourceTerm S  TQ ∈T TargetTerm (S)  (TP, TQ)  TRel+"
      by simp_all
    from reflTRel show "SP TQ. SP ∈S SourceTerm S  TQ ∈T TargetTerm (S)
                         (SP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
      by blast
  next
    case (source S)
    show "SP SQ. SP ∈S SourceTerm S  SQ ∈S SourceTerm S  SP = SQ"
      by simp
    show "SP TQ. SP ∈S SourceTerm S  TQ ∈T SourceTerm S
           (SP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
     and "TP SQ. TP ∈T SourceTerm S  SQ ∈S SourceTerm S  False"
     and "TP TQ. TP ∈T SourceTerm S  TQ ∈T SourceTerm S  (TP, TQ)  TRel+"
      by simp_all
  next
    case (target T1 T2)
    show "SP SQ. SP ∈S TargetTerm T1  SQ ∈S TargetTerm T2  SP = SQ"
     and "SP TQ. SP ∈S TargetTerm T1  TQ ∈T TargetTerm T2
           (SP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
     and "TP SQ. TP ∈T TargetTerm T1  SQ ∈S TargetTerm T2  False"
      by simp_all
    assume "(T1, T2)  TRel"
    thus "TP TQ. TP ∈T TargetTerm T1  TQ ∈T TargetTerm T2  (TP, TQ)  TRel+"
      by simp
  next
    case (trans P Q R)
    assume A1: "SP SQ. SP ∈S P  SQ ∈S Q  SP = SQ"
       and A2: "SP TQ. SP ∈S P  TQ ∈T Q
                 (SP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
       and A3: "TP SQ. TP ∈T P  SQ ∈S Q  False"
       and A4: "TP TQ. TP ∈T P  TQ ∈T Q  (TP, TQ)  TRel+"
       and A5: "SQ SR. SQ ∈S Q  SR ∈S R  SQ = SR"
       and A6: "SQ TR. SQ ∈S Q  TR ∈T R
                 (SQ, TR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
       and A7: "TQ SR. TQ ∈T Q  SR ∈S R  False"
       and A8: "TQ TR. TQ ∈T Q  TR ∈T R  (TQ, TR)  TRel+"
    show "SP SR. SP ∈S P  SR ∈S R  SP = SR"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A1 A5 show "SP SR. SP ∈S P  SR ∈S R  SP = SR"
        by blast
    next
      case (TargetTerm TQ)
      assume "TQ ∈T Q"
      with A7 show ?thesis
        by blast
    qed
    show "SP TR. SP ∈S P  TR ∈T R
           (SP, TR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A1 A6 show ?thesis
        by blast
    next
      case (TargetTerm TQ)
      assume A9: "TQ ∈T Q"
      show "SP TR. SP ∈S P  TR ∈T R
             (SP, TR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
      proof clarify
        fix SP TR
        assume "SP ∈S P"
        with A2 A9 have "(SP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
          by simp
        moreover assume "TR ∈T R"
        with A8 A9 have "(TQ, TR)  TRel+"
          by simp
        hence "(TQ, TR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
        proof induct
          fix T2
          assume "(TQ, T2)  TRel"
          thus "(TQ, T2)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
            by blast
        next
          case (step T2 T3)
          assume "(TQ, T2)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
          moreover assume "(T2, T3)  TRel"
          hence "(T2, T3)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
            by blast
          ultimately show "(TQ, T3)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
            by simp
        qed
        ultimately show "(SP, TR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
          by simp
      qed
    qed
    show "TP SR. TP ∈T P  SR ∈S R  False"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A3 show ?thesis
        by blast
    next
      case (TargetTerm TQ)
      assume "TQ ∈T Q"
      with A7 show ?thesis
        by blast
    qed
    show "TP TR. TP ∈T P  TR ∈T R  (TP, TR)  TRel+"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A3 show ?thesis
        by blast
    next
      case (TargetTerm TQ)
      assume "TQ ∈T Q"
      with A4 A8 show "TP TR. TP ∈T P  TR ∈T R  (TP, TR)  TRel+"
        by auto
    qed
  qed
qed

lemma (in encoding) indRelLTPO_to_TRel:
  fixes TRel :: "('procT × 'procT) set"
    and P Q  :: "('procS, 'procT) Proc"
  assumes rel: "P ≲⟦⋅⟧LT<TRel> Q"
  shows "SP SQ. SP ∈S P  SQ ∈S Q  SP = SQ"
    and "SP TQ. SP ∈S P  TQ ∈T Q  False"
    and "TP SQ. TP ∈T P  SQ ∈S Q
          (TP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
    and "TP TQ. TP ∈T P  TQ ∈T Q  (TP, TQ)  TRel+"
proof -
  have reflTRel: "S. (S, S)  TRel  {(T1, T2). S. T1 = S  T2 = S}"
    by auto
  from rel show "SP SQ. SP ∈S P  SQ ∈S Q  SP = SQ"
            and "SP TQ. SP ∈S P  TQ ∈T Q  False"
            and "TP SQ. TP ∈T P  SQ ∈S Q
                  (TP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
            and "TP TQ. TP ∈T P  TQ ∈T Q  (TP, TQ)  TRel+"
  proof induct
    case (encL S)
    show "SP SQ. SP ∈S TargetTerm (S)  SQ ∈S SourceTerm S   SP = SQ"
     and "SP TQ. SP ∈S TargetTerm (S)  TQ ∈T SourceTerm S  False"
     and "TP TQ. TP ∈T TargetTerm (S)  TQ ∈T SourceTerm S  (TP, TQ)  TRel+"
      by simp_all
    from reflTRel show "TP SQ. TP ∈T TargetTerm (S)  SQ ∈S SourceTerm S
                         (TP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
      by blast
  next
    case (source S)
    show "SP SQ. SP ∈S SourceTerm S  SQ ∈S SourceTerm S  SP = SQ"
      by simp
    show "SP TQ. SP ∈S SourceTerm S  TQ ∈T SourceTerm S  False"
     and "TP SQ. TP ∈T SourceTerm S  SQ ∈S SourceTerm S
           (TP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
     and "TP TQ. TP ∈T SourceTerm S  TQ ∈T SourceTerm S  (TP, TQ)  TRel+"
      by simp_all
  next
    case (target T1 T2)
    show "SP SQ. SP ∈S TargetTerm T1  SQ ∈S TargetTerm T2  SP = SQ"
     and "SP TQ. SP ∈S TargetTerm T1  TQ ∈T TargetTerm T2  False"
     and "TP SQ. TP ∈T TargetTerm T1  SQ ∈S TargetTerm T2
           (TP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
      by simp_all
    assume "(T1, T2)  TRel"
    thus "TP TQ. TP ∈T TargetTerm T1  TQ ∈T TargetTerm T2  (TP, TQ)  TRel+"
      by simp
  next
    case (trans P Q R)
    assume A1: "SP SQ. SP ∈S P  SQ ∈S Q  SP = SQ"
       and A2: "SP TQ. SP ∈S P  TQ ∈T Q  False"
       and A3: "TP SQ. TP ∈T P  SQ ∈S Q
                 (TP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
       and A4: "TP TQ. TP ∈T P  TQ ∈T Q  (TP, TQ)  TRel+"
       and A5: "SQ SR. SQ ∈S Q  SR ∈S R  SQ = SR"
       and A6: "SQ TR. SQ ∈S Q  TR ∈T R  False"
       and A7: "TQ SR. TQ ∈T Q  SR ∈S R
                 (TQ, SR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
       and A8: "TQ TR. TQ ∈T Q  TR ∈T R  (TQ, TR)  TRel+"
    show "SP SR. SP ∈S P  SR ∈S R  SP = SR"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A1 A5 show "SP SR. SP ∈S P  SR ∈S R  SP = SR"
        by blast
    next
      case (TargetTerm TQ)
      assume "TQ ∈T Q"
      with A2 show ?thesis
        by blast
    qed
    show "SP TR. SP ∈S P  TR ∈T R  False"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A6 show ?thesis
        by blast
    next
      case (TargetTerm TQ)
      assume "TQ ∈T Q"
      with A2 show ?thesis
        by blast
    qed
    show "TP SR. TP ∈T P  SR ∈S R
           (TP, SR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A3 A5 show "TP SR. TP ∈T P  SR ∈S R
                        (TP, SR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
        by blast
    next
      case (TargetTerm TQ)
      assume A9: "TQ ∈T Q"
      show "TP SR. TP ∈T P  SR ∈S R
             (TP, SR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
      proof clarify
        fix TP SR
        assume "TP ∈T P"
        with A4 A9 have "(TP, TQ)  TRel+"
          by simp
        hence "(TP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
        proof induct
          fix T2
          assume "(TP, T2)  TRel"
          thus "(TP, T2)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
            by blast
        next
          case (step T2 T3)
          assume "(TP, T2)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
          moreover assume "(T2, T3)  TRel"
          hence "(T2, T3)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
            by blast
          ultimately show "(TP, T3)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
            by simp
        qed
        moreover assume "SR ∈S R"
        with A7 A9 have "(TQ, SR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
          by simp
        ultimately show "(TP, SR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
          by simp
      qed
    qed
    show "TP TR. TP ∈T P  TR ∈T R  (TP, TR)  TRel+"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A6 show ?thesis
        by blast
    next
      case (TargetTerm TQ)
      assume "TQ ∈T Q"
      with A4 A8 show "TP TR. TP ∈T P  TR ∈T R  (TP, TR)  TRel+"
        by auto
    qed
  qed
qed

lemma (in encoding) indRelTEQ_to_TRel:
  fixes TRel :: "('procT × 'procT) set"
    and P Q  :: "('procS, 'procT) Proc"
  assumes rel: "P ∼⟦⋅⟧T<TRel> Q"
  shows "SP SQ. SP ∈S P  SQ ∈S Q
          (SP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
    and "SP TQ. SP ∈S P  TQ ∈T Q
          (SP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
    and "TP SQ. TP ∈T P  SQ ∈S Q
          (TP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
    and "TP TQ. TP ∈T P  TQ ∈T Q
          (TP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
proof -
  have reflTRel: "S. (S, S)  TRel  {(T1, T2). S. T1 = S  T2 = S}"
    by auto
  from rel show "SP SQ. SP ∈S P  SQ ∈S Q
                  (SP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
            and "SP TQ. SP ∈S P  TQ ∈T Q
                  (SP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
            and "TP SQ. TP ∈T P  SQ ∈S Q
                  (TP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
            and "TP TQ. TP ∈T P  TQ ∈T Q
                  (TP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
  proof induct
    case (encR S)
    show "SP SQ. SP ∈S SourceTerm S  SQ ∈S TargetTerm (S)
           (SP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
     and "TP SQ. TP ∈T SourceTerm S  SQ ∈S TargetTerm (S)
           (TP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
     and "TP TQ. TP ∈T SourceTerm S  TQ ∈T TargetTerm (S)
           (TP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
      by simp+
    from reflTRel show "SP TQ. SP ∈S SourceTerm S  TQ ∈T TargetTerm (S)
                         (SP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
      by blast
  next
    case (encL S)
    show "SP SQ. SP ∈S TargetTerm (S)  SQ ∈S SourceTerm S
           (SP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
     and "SP TQ. SP ∈S TargetTerm (S)  TQ ∈T SourceTerm S
           (SP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
     and "TP TQ. TP ∈T TargetTerm (S)  TQ ∈T SourceTerm S
           (TP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
      by simp+
    from reflTRel show "TP SQ. TP ∈T TargetTerm (S)  SQ ∈S SourceTerm S
                         (TP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
      by blast
  next
    case (target T1 T2)
    show "SP SQ. SP ∈S TargetTerm T1  SQ ∈S TargetTerm T2
           (SP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
     and "SP TQ. SP ∈S TargetTerm T1  TQ ∈T TargetTerm T2
           (SP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
     and "TP SQ. TP ∈T TargetTerm T1  SQ ∈S TargetTerm T2
           (TP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
      by simp+
    assume "(T1, T2)  TRel"
    thus "TP TQ. TP ∈T TargetTerm T1  TQ ∈T TargetTerm T2
           (TP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
      by blast
  next
    case (trans P Q R)
    assume A1: "SP SQ. SP ∈S P  SQ ∈S Q
                 (SP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
       and A2: "SP TQ. SP ∈S P  TQ ∈T Q
                 (SP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
       and A3: "TP SQ. TP ∈T P  SQ ∈S Q
                 (TP, SQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
       and A4: "TP TQ. TP ∈T P  TQ ∈T Q
                 (TP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
       and A5: "SQ SR. SQ ∈S Q  SR ∈S R
                 (SQ, SR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
       and A6: "SQ TR. SQ ∈S Q  TR ∈T R
                 (SQ, TR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
       and A7: "TQ SR. TQ ∈T Q  SR ∈S R
                 (TQ, SR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
       and A8: "TQ TR. TQ ∈T Q  TR ∈T R
                 (TQ, TR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
    show "SP SR. SP ∈S P  SR ∈S R
           (SP, SR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A1 A5 show ?thesis
        by auto
    next
      case (TargetTerm TQ)
      assume "TQ ∈T Q"
      with A2 A7 show ?thesis
        by auto
    qed
    show "SP TR. SP ∈S P  TR ∈T R  (SP, TR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A1 A6 show ?thesis
        by auto
    next
      case (TargetTerm TQ)
      assume "TQ ∈T Q"
      with A2 A8 show ?thesis
        by auto
    qed
    show "TP SR. TP ∈T P  SR ∈S R  (TP, SR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A3 A5 show ?thesis
        by auto
    next
      case (TargetTerm TQ)
      assume "TQ ∈T Q"
      with A4 A7 show ?thesis
        by auto
    qed
    show "TP TR. TP ∈T P  TR ∈T R  (TP, TR)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A3 A6 show ?thesis
        by auto
    next
      case (TargetTerm TQ)
      assume "TQ ∈T Q"
      with A4 A8 show ?thesis
        by auto
    qed
  qed
qed

lemma (in encoding) trans_closure_of_TRel_refl_cond:
  fixes TRel  :: "('procT × 'procT) set"
    and TP TQ :: "'procT"
  assumes "(TP, TQ)  (TRel  {(T1, T2). S. T1 = S  T2 = S})+"
  shows "(TP, TQ)  TRel*"
    using assms
proof induct
  fix TQ
  assume "(TP, TQ)  TRel  {(T1, T2). S. T1 = S  T2 = S}"
  thus "(TP, TQ)  TRel*"
    by auto
next
  case (step TQ TR)
  assume "(TP, TQ)  TRel*"
  moreover assume "(TQ, TR)  TRel  {(T1, T2). S. T1 = S  T2 = S}"
  hence "(TQ, TR)  TRel*"
    by blast
  ultimately show "(TP, TR)  TRel*"
    by simp
qed

text ‹Note that if indRelRTPO relates a source term S to a target term T, then the translation of
        S is equal to T or indRelRTPO also relates the translation of S to T.›

lemma (in encoding) indRelRTPO_relates_source_target:
  fixes TRel :: "('procT × 'procT) set"
    and S    :: "'procS"
    and T    :: "'procT"
  assumes pair: "SourceTerm S ≲⟦⋅⟧RT<TRel> TargetTerm T"
  shows "(TargetTerm (S), TargetTerm T)  (indRelRTPO TRel)="
proof -
  from pair have "(S, T)  TRel*"
      using indRelRTPO_to_TRel(2)[where TRel="TRel"] trans_closure_of_TRel_refl_cond
    by simp
  hence "S = T  (S, T)  TRel+"
      using rtrancl_eq_or_trancl[of "S" T TRel]
    by blast
  moreover have "S = T  (TargetTerm (S), TargetTerm T)  (indRelRTPO TRel)="
    by simp
  moreover
  have "(S, T)  TRel+  (TargetTerm (S), TargetTerm T)  (indRelRTPO TRel)="
      using transitive_closure_of_TRel_to_indRelRTPO[where TRel="TRel"]
    by simp
  ultimately show "(TargetTerm (S), TargetTerm T)  (indRelRTPO TRel)="
    by blast
qed

text ‹If indRelRTPO, indRelLTPO, or indRelTPO preserves barbs then so does the corresponding
        target term relation.›

lemma (in encoding_wrt_barbs) rel_with_target_impl_TRel_preserves_barbs:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes preservation: "rel_preserves_barbs Rel (STCalWB SWB TWB)"
      and targetInRel:  "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
  shows "rel_preserves_barbs TRel TWB"
proof clarify
  fix TP TQ a
  assume "(TP, TQ)  TRel"
  with targetInRel have "(TargetTerm TP, TargetTerm TQ)  Rel"
    by blast
  moreover assume "TP↓<TWB>a"
  hence "TargetTerm TP↓.a"
    by simp
  ultimately have "TargetTerm TQ↓.a"
      using preservation preservation_of_barbs_in_barbed_encoding[where Rel="Rel"]
    by blast
  thus "TQ↓<TWB>a"
    by simp
qed

lemma (in encoding_wrt_barbs) indRelRTPO_impl_TRel_preserves_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes preservation: "rel_preserves_barbs (indRelRTPO TRel) (STCalWB SWB TWB)"
  shows "rel_preserves_barbs TRel TWB"
      using preservation
            rel_with_target_impl_TRel_preserves_barbs[where Rel="indRelRTPO TRel" and TRel="TRel"]
    by (simp add: indRelRTPO.target)

lemma (in encoding_wrt_barbs) indRelLTPO_impl_TRel_preserves_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes preservation: "rel_preserves_barbs (indRelLTPO TRel) (STCalWB SWB TWB)"
  shows "rel_preserves_barbs TRel TWB"
      using preservation
            rel_with_target_impl_TRel_preserves_barbs[where Rel="indRelLTPO TRel" and TRel="TRel"]
    by (simp add: indRelLTPO.target)

lemma (in encoding_wrt_barbs) indRelTEQ_impl_TRel_preserves_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes preservation: "rel_preserves_barbs (indRelTEQ TRel) (STCalWB SWB TWB)"
  shows "rel_preserves_barbs TRel TWB"
      using preservation
            rel_with_target_impl_TRel_preserves_barbs[where Rel="indRelTEQ TRel" and TRel="TRel"]
    by (simp add: indRelTEQ.target)

lemma (in encoding_wrt_barbs) rel_with_target_impl_TRel_weakly_preserves_barbs:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes preservation: "rel_weakly_preserves_barbs Rel (STCalWB SWB TWB)"
      and targetInRel:  "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
  shows "rel_weakly_preserves_barbs TRel TWB"
proof clarify
  fix TP TQ a TP'
  assume "(TP, TQ)  TRel"
  with targetInRel have "(TargetTerm TP, TargetTerm TQ)  Rel"
    by blast
  moreover assume "TP (Calculus TWB)* TP'" and "TP'↓<TWB>a"
  hence "TargetTerm TP⇓.a"
    by blast
  ultimately have "TargetTerm TQ⇓.a"
      using preservation weak_preservation_of_barbs_in_barbed_encoding[where Rel="Rel"]
    by blast
  thus "TQ⇓<TWB>a"
    by simp
qed

lemma (in encoding_wrt_barbs) indRelRTPO_impl_TRel_weakly_preserves_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes preservation: "rel_weakly_preserves_barbs (indRelRTPO TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_preserves_barbs TRel TWB"
      using preservation rel_with_target_impl_TRel_weakly_preserves_barbs[where
                          Rel="indRelRTPO TRel" and TRel="TRel"]
    by (simp add: indRelRTPO.target)

lemma (in encoding_wrt_barbs) indRelLTPO_impl_TRel_weakly_preserves_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes preservation: "rel_weakly_preserves_barbs (indRelLTPO TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_preserves_barbs TRel TWB"
      using preservation rel_with_target_impl_TRel_weakly_preserves_barbs[where
                          Rel="indRelLTPO TRel" and TRel="TRel"]
    by (simp add: indRelLTPO.target)

lemma (in encoding_wrt_barbs) indRelTEQ_impl_TRel_weakly_preserves_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes preservation: "rel_weakly_preserves_barbs (indRelTEQ TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_preserves_barbs TRel TWB"
      using preservation rel_with_target_impl_TRel_weakly_preserves_barbs[where
                          Rel="indRelTEQ TRel" and TRel="TRel"]
    by (simp add: indRelTEQ.target)

text ‹If indRelRTPO, indRelLTPO, or indRelTPO reflects barbs then so does the corresponding
        target term relation.›

lemma (in encoding_wrt_barbs) rel_with_target_impl_TRel_reflects_barbs:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes reflection: "rel_reflects_barbs Rel (STCalWB SWB TWB)"
      and targetInRel:  "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
  shows "rel_reflects_barbs TRel TWB"
proof clarify
  fix TP TQ a
  assume "(TP, TQ)  TRel"
  with targetInRel have "(TargetTerm TP, TargetTerm TQ)  Rel"
    by blast
  moreover assume "TQ↓<TWB>a"
  hence "TargetTerm TQ↓.a"
    by simp
  ultimately have "TargetTerm TP↓.a"
      using reflection reflection_of_barbs_in_barbed_encoding[where Rel="Rel"]
    by blast
  thus "TP↓<TWB>a"
    by simp
qed

lemma (in encoding_wrt_barbs) indRelRTPO_impl_TRel_reflects_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes reflection: "rel_reflects_barbs (indRelRTPO TRel) (STCalWB SWB TWB)"
  shows "rel_reflects_barbs TRel TWB"
      using reflection
            rel_with_target_impl_TRel_reflects_barbs[where Rel="indRelRTPO TRel" and TRel="TRel"]
    by (simp add: indRelRTPO.target)

lemma (in encoding_wrt_barbs) indRelLTPO_impl_TRel_reflects_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes reflection: "rel_reflects_barbs (indRelLTPO TRel) (STCalWB SWB TWB)"
  shows "rel_reflects_barbs TRel TWB"
      using reflection
            rel_with_target_impl_TRel_reflects_barbs[where Rel="indRelLTPO TRel" and TRel="TRel"]
    by (simp add: indRelLTPO.target)

lemma (in encoding_wrt_barbs) indRelTEQ_impl_TRel_reflects_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes reflection: "rel_reflects_barbs (indRelTEQ TRel) (STCalWB SWB TWB)"
  shows "rel_reflects_barbs TRel TWB"
      using reflection
            rel_with_target_impl_TRel_reflects_barbs[where Rel="indRelTEQ TRel" and TRel="TRel"]
    by (simp add: indRelTEQ.target)

lemma (in encoding_wrt_barbs) rel_with_target_impl_TRel_weakly_reflects_barbs:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes reflection: "rel_weakly_reflects_barbs Rel (STCalWB SWB TWB)"
      and targetInRel:  "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
  shows "rel_weakly_reflects_barbs TRel TWB"
proof clarify
  fix TP TQ a TQ'
  assume "(TP, TQ)  TRel"
  with targetInRel have "(TargetTerm TP, TargetTerm TQ)  Rel"
    by blast
  moreover assume "TQ (Calculus TWB)* TQ'" and "TQ'↓<TWB>a"
  hence "TargetTerm TQ⇓.a"
    by blast
  ultimately have "TargetTerm TP⇓.a"
      using reflection weak_reflection_of_barbs_in_barbed_encoding[where Rel="Rel"]
    by blast
  thus "TP⇓<TWB>a"
    by simp
qed

lemma (in encoding_wrt_barbs) indRelRTPO_impl_TRel_weakly_reflects_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes reflection: "rel_weakly_reflects_barbs (indRelRTPO TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_reflects_barbs TRel TWB"
      using reflection rel_with_target_impl_TRel_weakly_reflects_barbs[where
                        Rel="indRelRTPO TRel" and TRel="TRel"]
    by (simp add: indRelRTPO.target)

lemma (in encoding_wrt_barbs) indRelLTPO_impl_TRel_weakly_reflects_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes reflection: "rel_weakly_reflects_barbs (indRelLTPO TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_reflects_barbs TRel TWB"
      using reflection rel_with_target_impl_TRel_weakly_reflects_barbs[where
                        Rel="indRelLTPO TRel" and TRel="TRel"]
    by (simp add: indRelLTPO.target)

lemma (in encoding_wrt_barbs) indRelTEQ_impl_TRel_weakly_reflects_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes reflection: "rel_weakly_reflects_barbs (indRelTEQ TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_reflects_barbs TRel TWB"
      using reflection rel_with_target_impl_TRel_weakly_reflects_barbs[where
                        Rel="indRelTEQ TRel" and TRel="TRel"]
    by (simp add: indRelTEQ.target)

text ‹If indRelRTPO, indRelLTPO, or indRelTPO respects barbs then so does the corresponding
        target term relation.›

lemma (in encoding_wrt_barbs) indRelRTPO_impl_TRel_respects_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes respection: "rel_respects_barbs (indRelRTPO TRel) (STCalWB SWB TWB)"
  shows "rel_respects_barbs TRel TWB"
      using respection indRelRTPO_impl_TRel_preserves_barbs[where TRel="TRel"]
            indRelRTPO_impl_TRel_reflects_barbs[where TRel="TRel"]
    by blast

lemma (in encoding_wrt_barbs) indRelLTPO_impl_TRel_respects_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes respection: "rel_respects_barbs (indRelLTPO TRel) (STCalWB SWB TWB)"
  shows "rel_respects_barbs TRel TWB"
      using respection indRelLTPO_impl_TRel_preserves_barbs[where TRel="TRel"]
            indRelLTPO_impl_TRel_reflects_barbs[where TRel="TRel"]
    by blast

lemma (in encoding_wrt_barbs) indRelTEQ_impl_TRel_respects_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes respection: "rel_respects_barbs (indRelTEQ TRel) (STCalWB SWB TWB)"
  shows "rel_respects_barbs TRel TWB"
      using respection indRelTEQ_impl_TRel_preserves_barbs[where TRel="TRel"]
            indRelTEQ_impl_TRel_reflects_barbs[where TRel="TRel"]
    by blast

lemma (in encoding_wrt_barbs) indRelRTPO_impl_TRel_weakly_respects_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes respection: "rel_weakly_respects_barbs (indRelRTPO TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_respects_barbs TRel TWB"
      using respection indRelRTPO_impl_TRel_weakly_preserves_barbs[where TRel="TRel"]
            indRelRTPO_impl_TRel_weakly_reflects_barbs[where TRel="TRel"]
    by blast

lemma (in encoding_wrt_barbs) indRelLTPO_impl_TRel_weakly_respects_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes respection: "rel_weakly_respects_barbs (indRelLTPO TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_respects_barbs TRel TWB"
      using respection indRelLTPO_impl_TRel_weakly_preserves_barbs[where TRel="TRel"]
            indRelLTPO_impl_TRel_weakly_reflects_barbs[where TRel="TRel"]
    by blast

lemma (in encoding_wrt_barbs) indRelTEQ_impl_TRel_weakly_respects_barbs:
  fixes TRel :: "('procT × 'procT) set"
  assumes respection: "rel_weakly_respects_barbs (indRelTEQ TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_respects_barbs TRel TWB"
      using respection indRelTEQ_impl_TRel_weakly_preserves_barbs[where TRel="TRel"]
            indRelTEQ_impl_TRel_weakly_reflects_barbs[where TRel="TRel"]
    by blast

text ‹If indRelRTPO, indRelLTPO, or indRelTEQ is a simulation then so is the corresponding target
        term relation.›

lemma (in encoding) rel_with_target_impl_transC_TRel_is_weak_reduction_simulation:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes sim:    "weak_reduction_simulation Rel (STCal Source Target)"
      and target: "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
      and trel:   "T1 T2. (TargetTerm T1, TargetTerm T2)  Rel  (T1, T2)  TRel+"
  shows "weak_reduction_simulation (TRel+) Target"
proof clarify
  fix TP TQ TP'
  assume "(TP, TQ)  TRel+" and "TP Target* TP'"
  thus "TQ'. TQ Target* TQ'  (TP', TQ')  TRel+"
  proof (induct arbitrary: TP')
    fix TQ TP'
    assume "(TP, TQ)  TRel"
    with target have "(TargetTerm TP, TargetTerm TQ)  Rel"
      by simp
    moreover assume "TP Target* TP'"
    hence "TargetTerm TP (STCal Source Target)* (TargetTerm TP')"
      by (simp add: STCal_steps)
    ultimately obtain Q' where A2: "TargetTerm TQ (STCal Source Target)* Q'"
                           and A3: "(TargetTerm TP', Q')  Rel"
        using sim
      by blast
    from A2 obtain TQ' where A4: "TQ Target* TQ'" and A5: "TQ' ∈T Q'"
      by (auto simp add: STCal_steps)
    from A3 A5 trel have "(TP', TQ')  TRel+"
      by simp
    with A4 show "TQ'. TQ Target* TQ'  (TP', TQ')  TRel+"
      by blast
  next
    case (step TQ TR)
    assume "TP Target* TP'"
       and "TP'. TP Target* TP'  TQ'. TQ Target* TQ'  (TP', TQ')  TRel+"
    from this obtain TQ' where B1: "TQ Target* TQ'" and B2: "(TP', TQ')  TRel+"
      by blast
    assume "(TQ, TR)  TRel"
    with target have "(TargetTerm TQ, TargetTerm TR)  Rel"
      by simp
    moreover from B1 have "TargetTerm TQ (STCal Source Target)* (TargetTerm TQ')"
      by (simp add: STCal_steps)
    ultimately obtain R' where B3: "TargetTerm TR (STCal Source Target)* R'"
                           and B4: "(TargetTerm TQ', R')  Rel"
        using sim
      by blast
    from B3 obtain TR' where B5: "TR' ∈T R'" and B6: "TR Target* TR'"
      by (auto simp add: STCal_steps)
    from B4 B5 trel have "(TQ', TR')  TRel+"
      by simp
    with B2 have "(TP', TR')  TRel+"
      by simp
    with B6 show "TR'. TR Target* TR'  (TP', TR')  TRel+"
      by blast
  qed
qed

lemma (in encoding) indRelRTPO_impl_TRel_is_weak_reduction_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "weak_reduction_simulation (indRelRTPO TRel) (STCal Source Target)"
  shows "weak_reduction_simulation (TRel+) Target"
      using sim indRelRTPO.target[where TRel="TRel"] indRelRTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_weak_reduction_simulation[where
             Rel="indRelRTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding) indRelLTPO_impl_TRel_is_weak_reduction_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "weak_reduction_simulation (indRelLTPO TRel) (STCal Source Target)"
  shows "weak_reduction_simulation (TRel+) Target"
      using sim indRelLTPO.target[where TRel="TRel"] indRelLTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_weak_reduction_simulation[where
             Rel="indRelLTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding) rel_with_target_impl_transC_TRel_is_weak_reduction_simulation_rev:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes sim:    "weak_reduction_simulation (Rel¯) (STCal Source Target)"
      and target: "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
      and trel:   "T1 T2. (TargetTerm T1, TargetTerm T2)  Rel  (T1, T2)  TRel+"
  shows "weak_reduction_simulation ((TRel+)¯) Target"
proof clarify
  fix TP TQ TP'
  assume "(TQ, TP)  TRel+"
  moreover assume "TP Target* TP'"
  ultimately show "TQ'. TQ Target* TQ'  (TP', TQ')  (TRel+)¯"
  proof (induct arbitrary: TP')
    fix TP TP'
    assume "(TQ, TP)  TRel"
    with target have "(TargetTerm TP, TargetTerm TQ)  Rel¯"
      by simp
    moreover assume "TP Target* TP'"
    hence "TargetTerm TP (STCal Source Target)* (TargetTerm TP')"
      by (simp add: STCal_steps)
    ultimately obtain Q' where A2: "TargetTerm TQ (STCal Source Target)* Q'"
                           and A3: "(TargetTerm TP', Q')  Rel¯"
        using sim
      by blast
    from A2 obtain TQ' where A4: "TQ Target* TQ'" and A5: "TQ' ∈T Q'"
      by (auto simp add: STCal_steps(2))
    from A3 A5 trel have "(TP', TQ')  (TRel+)¯"
      by simp
    with A4 show "TQ'. TQ Target* TQ'  (TP', TQ')  (TRel+)¯"
      by blast
  next
    case (step TR TP TP')
    assume "TP Target* TP'"
    hence "TargetTerm TP (STCal Source Target)* (TargetTerm TP')"
      by (simp add: STCal_steps)
    moreover assume "(TR, TP)  TRel"
    with target have "(TargetTerm TP, TargetTerm TR)  Rel¯"
      by simp
    ultimately obtain R' where B1: "TargetTerm TR (STCal Source Target)* R'"
                           and B2: "(TargetTerm TP', R')  Rel¯"
        using sim
      by blast
    from B1 obtain TR' where B3: "TR' ∈T R'" and B4: "TR Target* TR'"
      by (auto simp add: STCal_steps)
    assume "TR'. TR Target* TR'  TQ'. TQ Target* TQ'  (TR', TQ')  (TRel+)¯"
    with B4 obtain TQ' where B5: "TQ Target* TQ'" and B6: "(TR', TQ')  (TRel+)¯"
      by blast
    from B6 have "(TQ', TR')  TRel+"
      by simp
    moreover from B2 B3 trel have "(TR', TP')  TRel+"
      by simp
    ultimately have "(TP', TQ')  (TRel+)¯"
      by simp
    with B5 show "TQ'. TQ Target* TQ'  (TP', TQ')  (TRel+)¯"
      by blast
  qed
qed

lemma (in encoding) indRelRTPO_impl_TRel_is_weak_reduction_simulation_rev:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "weak_reduction_simulation ((indRelRTPO TRel)¯) (STCal Source Target)"
  shows "weak_reduction_simulation ((TRel+)¯) Target"
      using sim indRelRTPO.target[where TRel="TRel"] indRelRTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_weak_reduction_simulation_rev[where
             Rel="indRelRTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding) indRelLTPO_impl_TRel_is_weak_reduction_simulation_rev:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "weak_reduction_simulation ((indRelLTPO TRel)¯) (STCal Source Target)"
  shows "weak_reduction_simulation ((TRel+)¯) Target"
      using sim indRelLTPO.target[where TRel="TRel"] indRelLTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_weak_reduction_simulation_rev[where
             Rel="indRelLTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding) rel_with_target_impl_reflC_transC_TRel_is_weak_reduction_simulation:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes sim:    "weak_reduction_simulation Rel (STCal Source Target)"
      and target: "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
      and trel:   "T1 T2. (TargetTerm T1, TargetTerm T2)  Rel  (T1, T2)  TRel*"
  shows "weak_reduction_simulation (TRel*) Target"
proof clarify
  fix TP TQ TP'
  assume "(TP, TQ)  TRel*" and "TP Target* TP'"
  thus "TQ'. TQ Target* TQ'  (TP', TQ')  TRel*"
  proof (induct arbitrary: TP')
    fix TP'
    assume "TP Target* TP'"
    moreover have "(TP', TP')  TRel*"
      by simp
    ultimately show "TQ'. TP Target* TQ'  (TP', TQ')  TRel*"
      by blast
  next
    case (step TQ TR)
    assume "TP Target* TP'"
       and "TP'. TP Target* TP'  TQ'. TQ Target* TQ'  (TP', TQ')  TRel*"
    from this obtain TQ' where B1: "TQ Target* TQ'" and B2: "(TP', TQ')  TRel*"
      by blast
    assume "(TQ, TR)  TRel"
    with target have "(TargetTerm TQ, TargetTerm TR)  Rel"
      by simp
    moreover from B1 have "TargetTerm TQ (STCal Source Target)* (TargetTerm TQ')"
      by (simp add: STCal_steps)
    ultimately obtain R' where B3: "TargetTerm TR (STCal Source Target)* R'"
                           and B4: "(TargetTerm TQ', R')  Rel"
        using sim
      by blast
    from B3 obtain TR' where B5: "TR' ∈T R'" and B6: "TR Target* TR'"
      by (auto simp add: STCal_steps)
    from B4 B5 trel have "(TQ', TR')  TRel*"
      by simp
    with B2 have "(TP', TR')  TRel*"
      by simp
    with B6 show "TR'. TR Target* TR'  (TP', TR')  TRel*"
      by blast
  qed
qed

lemma (in encoding) indRelTEQ_impl_TRel_is_weak_reduction_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "weak_reduction_simulation (indRelTEQ TRel) (STCal Source Target)"
  shows "weak_reduction_simulation (TRel*) Target"
      using sim indRelTEQ.target[where TRel="TRel"] indRelTEQ_to_TRel(4)[where TRel="TRel"]
            trans_closure_of_TRel_refl_cond
            rel_with_target_impl_reflC_transC_TRel_is_weak_reduction_simulation[where
             Rel="indRelTEQ TRel" and TRel="TRel"]
    by blast

lemma (in encoding) rel_with_target_impl_transC_TRel_is_strong_reduction_simulation:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes sim:    "strong_reduction_simulation Rel (STCal Source Target)"
      and target: "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
      and trel:   "T1 T2. (TargetTerm T1, TargetTerm T2)  Rel  (T1, T2)  TRel+"
  shows "strong_reduction_simulation (TRel+) Target"
proof clarify
  fix TP TQ TP'
  assume "(TP, TQ)  TRel+" and "TP Target TP'"
  thus "TQ'. TQ Target TQ'  (TP', TQ')  TRel+"
  proof (induct arbitrary: TP')
    fix TQ TP'
    assume "(TP, TQ)  TRel"
    with target have "(TargetTerm TP, TargetTerm TQ)  Rel"
      by simp
    moreover assume "TP Target TP'"
    hence "TargetTerm TP (STCal Source Target) (TargetTerm TP')"
      by (simp add: STCal_step)
    ultimately obtain Q' where A2: "TargetTerm TQ (STCal Source Target) Q'"
                           and A3: "(TargetTerm TP', Q')  Rel"
        using sim
      by blast
    from A2 obtain TQ' where A4: "TQ Target TQ'" and A5: "TQ' ∈T Q'"
      by (auto simp add: STCal_step)
    from A3 A5 trel have "(TP', TQ')  TRel+"
      by simp
    with A4 show "TQ'. TQ Target TQ'  (TP', TQ')  TRel+"
      by blast
  next
    case (step TQ TR)
    assume "TP Target TP'"
       and "TP'. TP Target TP'  TQ'. TQ Target TQ'  (TP', TQ')  TRel+"
    from this obtain TQ' where B1: "TQ Target TQ'" and B2: "(TP', TQ')  TRel+"
      by blast
    assume "(TQ, TR)  TRel"
    with target have "(TargetTerm TQ, TargetTerm TR)  Rel"
      by simp
    moreover from B1 have "TargetTerm TQ (STCal Source Target) (TargetTerm TQ')"
      by (simp add: STCal_step)
    ultimately obtain R' where B3: "TargetTerm TR (STCal Source Target) R'"
                           and B4: "(TargetTerm TQ', R')  Rel"
        using sim
      by blast
    from B3 obtain TR' where B5: "TR' ∈T R'" and B6: "TR Target TR'"
      by (auto simp add: STCal_step)
    from B4 B5 trel have "(TQ', TR')  TRel+"
      by simp
    with B2 have "(TP', TR')  TRel+"
      by simp
    with B6 show "TR'. TR Target TR'  (TP', TR')  TRel+"
      by blast
  qed
qed

lemma (in encoding) indRelRTPO_impl_TRel_is_strong_reduction_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "strong_reduction_simulation (indRelRTPO TRel) (STCal Source Target)"
  shows "strong_reduction_simulation (TRel+) Target"
      using sim indRelRTPO.target[where TRel="TRel"] indRelRTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_strong_reduction_simulation[where
             Rel="indRelRTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding) indRelLTPO_impl_TRel_is_strong_reduction_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "strong_reduction_simulation (indRelLTPO TRel) (STCal Source Target)"
  shows "strong_reduction_simulation (TRel+) Target"
      using sim indRelLTPO.target[where TRel="TRel"] indRelLTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_strong_reduction_simulation[where
             Rel="indRelLTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding) rel_with_target_impl_transC_TRel_is_strong_reduction_simulation_rev:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes sim:    "strong_reduction_simulation (Rel¯) (STCal Source Target)"
      and target: "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
      and trel:   "T1 T2. (TargetTerm T1, TargetTerm T2)  Rel  (T1, T2)  TRel+"
  shows "strong_reduction_simulation ((TRel+)¯) Target"
proof clarify
  fix TP TQ TP'
  assume "(TQ, TP)  TRel+"
  moreover assume "TP Target TP'"
  ultimately show "TQ'. TQ Target TQ'  (TP', TQ')  (TRel+)¯"
  proof (induct arbitrary: TP')
    fix TP TP'
    assume "(TQ, TP)  TRel"
    with target have "(TargetTerm TP, TargetTerm TQ)  Rel¯"
      by simp
    moreover assume "TP Target TP'"
    hence "TargetTerm TP (STCal Source Target) (TargetTerm TP')"
      by (simp add: STCal_step)
    ultimately obtain Q' where A2: "TargetTerm TQ (STCal Source Target) Q'"
                           and A3: "(TargetTerm TP', Q')  Rel¯"
        using sim
      by blast
    from A2 obtain TQ' where A4: "TQ Target TQ'" and A5: "TQ' ∈T Q'"
      by (auto simp add: STCal_step(2))
    from A3 A5 trel have "(TP', TQ')  (TRel+)¯"
      by simp
    with A4 show "TQ'. TQ Target TQ'  (TP', TQ')  (TRel+)¯"
      by blast
  next
    case (step TP TR TR')
    assume "(TP, TR)  TRel"
    with target have "(TargetTerm TP, TargetTerm TR)  Rel"
      by simp
    moreover assume "TR Target TR'"
    hence "TargetTerm TR (STCal Source Target) (TargetTerm TR')"
      by (simp add: STCal_step)
    ultimately obtain P' where B1: "TargetTerm TP (STCal Source Target) P'"
                           and B2: "(P', TargetTerm TR')  Rel"
        using sim
      by blast
    from B1 obtain TP' where B3: "TP' ∈T P'" and B4: "TP Target TP'"
      by (auto simp add: STCal_step)
    assume "TP'. TP Target TP'  TQ'. TQ Target TQ'  (TP', TQ')  (TRel+)¯"
    with B4 obtain TQ' where B5: "TQ Target TQ'" and B6: "(TP', TQ')  (TRel+)¯"
      by blast
    from B2 B3 trel have "(TP', TR')  TRel+"
      by simp
    with B6 have "(TR', TQ')  (TRel+)¯"
      by simp
    with B5 show "TQ'. TQ Target TQ'  (TR', TQ')  (TRel+)¯"
      by blast
  qed
qed

lemma (in encoding) indRelRTPO_impl_TRel_is_strong_reduction_simulation_rev:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "strong_reduction_simulation ((indRelRTPO TRel)¯) (STCal Source Target)"
  shows "strong_reduction_simulation ((TRel+)¯) Target"
      using sim indRelRTPO.target[where TRel="TRel"] indRelRTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_strong_reduction_simulation_rev[where
             Rel="indRelRTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding) indRelLTPO_impl_TRel_is_strong_reduction_simulation_rev:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "strong_reduction_simulation ((indRelLTPO TRel)¯) (STCal Source Target)"
  shows "strong_reduction_simulation ((TRel+)¯) Target"
      using sim indRelLTPO.target[where TRel="TRel"] indRelLTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_strong_reduction_simulation_rev[where
             Rel="indRelLTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding) rel_with_target_impl_reflC_transC_TRel_is_strong_reduction_simulation:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes sim:    "strong_reduction_simulation Rel (STCal Source Target)"
      and target: "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
      and trel:   "T1 T2. (TargetTerm T1, TargetTerm T2)  Rel
                    (T1, T2)  TRel*"
  shows "strong_reduction_simulation (TRel*) Target"
proof clarify
  fix TP TQ TP'
  assume "(TP, TQ)  TRel*" and "TP Target TP'"
  thus "TQ'. TQ Target TQ'  (TP', TQ')  TRel*"
  proof (induct arbitrary: TP')
    fix TP'
    assume "TP Target TP'"
    moreover have "(TP', TP')  TRel*"
      by simp
    ultimately show "TQ'. TP Target TQ'  (TP', TQ')  TRel*"
      by blast
  next
    case (step TQ TR TP')
    assume "TP Target TP'"
       and "TP'. TP Target TP'  TQ'. TQ Target TQ'  (TP', TQ')  TRel*"
    from this obtain TQ' where B1: "TQ Target TQ'" and B2: "(TP', TQ')  TRel*"
      by blast
    assume "(TQ, TR)  TRel"
    with target have "(TargetTerm TQ, TargetTerm TR)  Rel"
      by simp
    moreover from B1 have "TargetTerm TQ (STCal Source Target) (TargetTerm TQ')"
      by (simp add: STCal_step)
    ultimately obtain R' where B3: "TargetTerm TR (STCal Source Target) R'"
                           and B4: "(TargetTerm TQ', R')  Rel"
        using sim
      by blast
    from B3 obtain TR' where B5: "TR' ∈T R'" and B6: "TR Target TR'"
      by (auto simp add: STCal_step)
    from B4 B5 trel have "(TQ', TR')  TRel*"
      by simp
    with B2 have "(TP', TR')  TRel*"
      by simp
    with B6 show "TR'. TR Target TR'  (TP', TR')  TRel*"
      by blast
  qed
qed

lemma (in encoding) indRelTEQ_impl_TRel_is_strong_reduction_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "strong_reduction_simulation (indRelTEQ TRel) (STCal Source Target)"
  shows "strong_reduction_simulation (TRel*) Target"
      using sim indRelTEQ.target[where TRel="TRel"] indRelTEQ_to_TRel(4)[where TRel="TRel"]
            trans_closure_of_TRel_refl_cond
            rel_with_target_impl_reflC_transC_TRel_is_strong_reduction_simulation[where
             Rel="indRelTEQ TRel" and TRel="TRel"]
    by blast

lemma (in encoding_wrt_barbs) indRelRTPO_impl_TRel_is_weak_barbed_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "weak_barbed_simulation (indRelRTPO TRel) (STCalWB SWB TWB)"
  shows "weak_barbed_simulation (TRel+) TWB"
proof
  from sim show "weak_reduction_simulation (TRel+) (Calculus TWB)"
      using indRelRTPO_impl_TRel_is_weak_reduction_simulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from sim show "rel_weakly_preserves_barbs (TRel+) TWB"
      using indRelRTPO_impl_TRel_weakly_preserves_barbs[where TRel="TRel"]
            weak_preservation_of_barbs_and_closures(2)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelLTPO_impl_TRel_is_weak_barbed_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "weak_barbed_simulation (indRelLTPO TRel) (STCalWB SWB TWB)"
  shows "weak_barbed_simulation (TRel+) TWB"
proof
  from sim show "weak_reduction_simulation (TRel+) (Calculus TWB)"
      using indRelLTPO_impl_TRel_is_weak_reduction_simulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from sim show "rel_weakly_preserves_barbs (TRel+) TWB"
      using indRelLTPO_impl_TRel_weakly_preserves_barbs[where TRel="TRel"]
            weak_preservation_of_barbs_and_closures(2)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelTEQ_impl_TRel_is_weak_barbed_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "weak_barbed_simulation (indRelTEQ TRel) (STCalWB SWB TWB)"
  shows "weak_barbed_simulation (TRel*) TWB"
proof
  from sim show "weak_reduction_simulation (TRel*) (Calculus TWB)"
      using indRelTEQ_impl_TRel_is_weak_reduction_simulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from sim show "rel_weakly_preserves_barbs (TRel*) TWB"
      using indRelTEQ_impl_TRel_weakly_preserves_barbs[where TRel="TRel"]
            weak_preservation_of_barbs_and_closures(3)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelRTPO_impl_TRel_is_strong_barbed_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "strong_barbed_simulation (indRelRTPO TRel) (STCalWB SWB TWB)"
  shows "strong_barbed_simulation (TRel+) TWB"
proof
  from sim show "strong_reduction_simulation (TRel+) (Calculus TWB)"
      using indRelRTPO_impl_TRel_is_strong_reduction_simulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from sim show "rel_preserves_barbs (TRel+) TWB"
      using indRelRTPO_impl_TRel_preserves_barbs[where TRel="TRel"]
            preservation_of_barbs_and_closures(2)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelLTPO_impl_TRel_is_strong_barbed_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "strong_barbed_simulation (indRelLTPO TRel) (STCalWB SWB TWB)"
  shows "strong_barbed_simulation (TRel+) TWB"
proof
  from sim refl show "strong_reduction_simulation (TRel+) (Calculus TWB)"
      using indRelLTPO_impl_TRel_is_strong_reduction_simulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from sim show "rel_preserves_barbs (TRel+) TWB"
      using indRelLTPO_impl_TRel_preserves_barbs[where TRel="TRel"]
            preservation_of_barbs_and_closures(2)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelTEQ_impl_TRel_is_strong_barbed_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes sim: "strong_barbed_simulation (indRelTEQ TRel) (STCalWB SWB TWB)"
  shows "strong_barbed_simulation (TRel*) TWB"
proof
  from sim refl show "strong_reduction_simulation (TRel*) (Calculus TWB)"
      using indRelTEQ_impl_TRel_is_strong_reduction_simulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from sim show "rel_preserves_barbs (TRel*) TWB"
      using indRelTEQ_impl_TRel_preserves_barbs[where TRel="TRel"]
            preservation_of_barbs_and_closures(3)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

text ‹If indRelRTPO, indRelLTPO, or indRelTEQ is a contrasimulation then so is the corresponding
        target term relation.›

lemma (in encoding) rel_with_target_impl_transC_TRel_is_weak_reduction_contrasimulation:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes conSim: "weak_reduction_contrasimulation Rel (STCal Source Target)"
      and target: "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
      and trel:   "T1 T2. (TargetTerm T1, TargetTerm T2)  Rel  (T1, T2)  TRel+"
  shows "weak_reduction_contrasimulation (TRel+) Target"
proof clarify
  fix TP TQ TP'
  assume "(TP, TQ)  TRel+" and "TP Target* TP'"
  thus "TQ'. TQ Target* TQ'  (TQ', TP')  TRel+"
  proof (induct arbitrary: TP')
    fix TQ TP'
    assume "(TP, TQ)  TRel"
    with target have "(TargetTerm TP, TargetTerm TQ)  Rel"
      by simp
    moreover assume "TP Target* TP'"
    hence "TargetTerm TP (STCal Source Target)* (TargetTerm TP')"
      by (simp add: STCal_steps)
    ultimately obtain Q' where A2: "TargetTerm TQ (STCal Source Target)* Q'"
                           and A3: "(Q', TargetTerm TP')  Rel"
        using conSim
      by blast
    from A2 obtain TQ' where A4: "TQ Target* TQ'" and A5: "TQ' ∈T Q'"
      by (auto simp add: STCal_steps)
    from A3 A5 trel have "(TQ', TP')  TRel+"
      by simp
    with A4 show "TQ'. TQ Target* TQ'  (TQ', TP')  TRel+"
      by blast
  next
    case (step TQ TR)
    assume "TP Target* TP'"
       and "TP'. TP Target* TP'  TQ'. TQ Target* TQ'  (TQ', TP')  TRel+"
    from this obtain TQ' where B1: "TQ Target* TQ'" and B2: "(TQ', TP')  TRel+"
      by blast
    assume "(TQ, TR)  TRel"
    with target have "(TargetTerm TQ, TargetTerm TR)  Rel"
      by simp
    moreover from B1 have "TargetTerm TQ (STCal Source Target)* (TargetTerm TQ')"
      by (simp add: STCal_steps)
    ultimately obtain R' where B3: "TargetTerm TR (STCal Source Target)* R'"
                           and B4: "(R', TargetTerm TQ')  Rel"
        using conSim
      by blast
    from B3 obtain TR' where B5: "TR' ∈T R'" and B6: "TR Target* TR'"
      by (auto simp add: STCal_steps)
    from B4 B5 trel have "(TR', TQ')  TRel+"
      by simp
    from this B2 have "(TR', TP')  TRel+"
      by simp
    with B6 show "TR'. TR Target* TR'  (TR', TP')  TRel+"
      by blast
  qed
qed

lemma (in encoding) indRelRTPO_impl_TRel_is_weak_reduction_contrasimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes conSim: "weak_reduction_contrasimulation (indRelRTPO TRel) (STCal Source Target)"
  shows "weak_reduction_contrasimulation (TRel+) Target"
      using conSim indRelRTPO.target[where TRel="TRel"] indRelRTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_weak_reduction_contrasimulation[where
             Rel="indRelRTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding) indRelLTPO_impl_TRel_is_weak_reduction_contrasimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes conSim: "weak_reduction_contrasimulation (indRelLTPO TRel) (STCal Source Target)"
  shows "weak_reduction_contrasimulation (TRel+) Target"
      using conSim indRelLTPO.target[where TRel="TRel"] indRelLTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_weak_reduction_contrasimulation[where
             Rel="indRelLTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding) rel_with_target_impl_reflC_transC_TRel_is_weak_reduction_contrasimulation:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes conSim: "weak_reduction_contrasimulation Rel (STCal Source Target)"
      and target: "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
      and trel:   "T1 T2. (TargetTerm T1, TargetTerm T2)  Rel  (T1, T2)  TRel*"
  shows "weak_reduction_contrasimulation (TRel*) Target"
proof clarify
  fix TP TQ TP'
  assume "(TP, TQ)  TRel*" and "TP Target* TP'"
  thus "TQ'. TQ Target* TQ'  (TQ', TP')  TRel*"
  proof (induct arbitrary: TP')
    fix TP'
    assume "TP Target* TP'"
    moreover have "(TP', TP')  TRel*"
      by simp
    ultimately show "TQ'. TP Target* TQ'  (TQ', TP')  TRel*"
      by blast
  next
    case (step TQ TR)
    assume "TP Target* TP'"
       and "TP'. TP Target* TP'  TQ'. TQ Target* TQ'  (TQ', TP')  TRel*"
    from this obtain TQ' where B1: "TQ Target* TQ'" and B2: "(TQ', TP')  TRel*"
      by blast
    assume "(TQ, TR)  TRel"
    with target have "(TargetTerm TQ, TargetTerm TR)  Rel"
      by simp
    moreover from B1 have "TargetTerm TQ (STCal Source Target)* (TargetTerm TQ')"
      by (simp add: STCal_steps)
    ultimately obtain R' where B3: "TargetTerm TR (STCal Source Target)* R'"
                           and B4: "(R', TargetTerm TQ')  Rel"
        using conSim
      by blast
    from B3 obtain TR' where B5: "TR' ∈T R'" and B6: "TR Target* TR'"
      by (auto simp add: STCal_steps)
    from B4 B5 trel have "(TR', TQ')  TRel*"
      by simp
    from this B2 have "(TR', TP')  TRel*"
      by simp
    with B6 show "TR'. TR Target* TR'  (TR', TP')  TRel*"
      by blast
  qed
qed

lemma (in encoding) indRelTEQ_impl_TRel_is_weak_reduction_contrasimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes conSim: "weak_reduction_contrasimulation (indRelTEQ TRel) (STCal Source Target)"
  shows "weak_reduction_contrasimulation (TRel*) Target"
      using conSim indRelTEQ.target[where TRel="TRel"] indRelTEQ_to_TRel(4)[where TRel="TRel"]
            trans_closure_of_TRel_refl_cond
            rel_with_target_impl_reflC_transC_TRel_is_weak_reduction_contrasimulation[where
             Rel="indRelTEQ TRel" and TRel="TRel"]
    by blast

lemma (in encoding_wrt_barbs) indRelRTPO_impl_TRel_is_weak_barbed_contrasimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes conSim: "weak_barbed_contrasimulation (indRelRTPO TRel) (STCalWB SWB TWB)"
  shows "weak_barbed_contrasimulation (TRel+) TWB"
proof
  from conSim show "weak_reduction_contrasimulation (TRel+) (Calculus TWB)"
      using indRelRTPO_impl_TRel_is_weak_reduction_contrasimulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from conSim show "rel_weakly_preserves_barbs (TRel+) TWB"
      using indRelRTPO_impl_TRel_weakly_preserves_barbs[where TRel="TRel"]
            weak_preservation_of_barbs_and_closures(2)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelLTPO_impl_TRel_is_weak_barbed_contrasimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes conSim: "weak_barbed_contrasimulation (indRelLTPO TRel) (STCalWB SWB TWB)"
  shows "weak_barbed_contrasimulation (TRel+) TWB"
proof
  from conSim show "weak_reduction_contrasimulation (TRel+) (Calculus TWB)"
      using indRelLTPO_impl_TRel_is_weak_reduction_contrasimulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from conSim show "rel_weakly_preserves_barbs (TRel+) TWB"
      using indRelLTPO_impl_TRel_weakly_preserves_barbs[where TRel="TRel"]
            weak_preservation_of_barbs_and_closures(2)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelTEQ_impl_TRel_is_weak_barbed_contrasimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes conSim: "weak_barbed_contrasimulation (indRelTEQ TRel) (STCalWB SWB TWB)"
  shows "weak_barbed_contrasimulation (TRel*) TWB"
proof
  from conSim show "weak_reduction_contrasimulation (TRel*) (Calculus TWB)"
      using indRelTEQ_impl_TRel_is_weak_reduction_contrasimulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from conSim show "rel_weakly_preserves_barbs (TRel*) TWB"
      using indRelTEQ_impl_TRel_weakly_preserves_barbs[where TRel="TRel"]
            weak_preservation_of_barbs_and_closures(3)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

text ‹If indRelRTPO, indRelLTPO, or indRelTEQ is a coupled simulation then so is the
        corresponding target term relation.›

lemma (in encoding) indRelRTPO_impl_TRel_is_weak_reduction_coupled_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes couSim: "weak_reduction_coupled_simulation (indRelRTPO TRel) (STCal Source Target)"
  shows "weak_reduction_coupled_simulation (TRel+) Target"
      using couSim weak_reduction_coupled_simulation_versus_simulation_and_contrasimulation
            refl indRelRTPO_impl_TRel_is_weak_reduction_simulation[where TRel="TRel"]
            indRelRTPO_impl_TRel_is_weak_reduction_contrasimulation[where TRel="TRel"]
    by blast

lemma (in encoding) indRelLTPO_impl_TRel_is_weak_reduction_coupled_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes couSim: "weak_reduction_coupled_simulation (indRelLTPO TRel) (STCal Source Target)"
  shows "weak_reduction_coupled_simulation (TRel+) Target"
      using couSim weak_reduction_coupled_simulation_versus_simulation_and_contrasimulation
            refl indRelLTPO_impl_TRel_is_weak_reduction_simulation[where TRel="TRel"]
            indRelLTPO_impl_TRel_is_weak_reduction_contrasimulation[where TRel="TRel"]
    by blast

lemma (in encoding) indRelTEQ_impl_TRel_is_weak_reduction_coupled_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes couSim: "weak_reduction_coupled_simulation (indRelTEQ TRel) (STCal Source Target)"
  shows "weak_reduction_coupled_simulation (TRel*) Target"
      using couSim weak_reduction_coupled_simulation_versus_simulation_and_contrasimulation
            refl indRelTEQ_impl_TRel_is_weak_reduction_simulation[where TRel="TRel"]
            indRelTEQ_impl_TRel_is_weak_reduction_contrasimulation[where TRel="TRel"]
    by blast

lemma (in encoding_wrt_barbs) indRelRTPO_impl_TRel_is_weak_barbed_coupled_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes couSim: "weak_barbed_coupled_simulation (indRelRTPO TRel) (STCalWB SWB TWB)"
  shows "weak_barbed_coupled_simulation (TRel+) TWB"
      using couSim weak_barbed_coupled_simulation_versus_simulation_and_contrasimulation
            refl indRelRTPO_impl_TRel_is_weak_barbed_simulation[where TRel="TRel"]
            indRelRTPO_impl_TRel_is_weak_barbed_contrasimulation[where TRel="TRel"]
    by blast

lemma (in encoding_wrt_barbs) indRelLTPO_impl_TRel_is_weak_barbed_coupled_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes couSim: "weak_barbed_coupled_simulation (indRelLTPO TRel) (STCalWB SWB TWB)"
  shows "weak_barbed_coupled_simulation (TRel+) TWB"
      using couSim weak_barbed_coupled_simulation_versus_simulation_and_contrasimulation
            refl indRelLTPO_impl_TRel_is_weak_barbed_simulation[where TRel="TRel"]
            indRelLTPO_impl_TRel_is_weak_barbed_contrasimulation[where TRel="TRel"]
    by blast

lemma (in encoding_wrt_barbs) indRelTEQ_impl_TRel_is_weak_barbed_coupled_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes couSim: "weak_barbed_coupled_simulation (indRelTEQ TRel) (STCalWB SWB TWB)"
  shows "weak_barbed_coupled_simulation (TRel*) TWB"
      using couSim weak_barbed_coupled_simulation_versus_simulation_and_contrasimulation
            refl indRelTEQ_impl_TRel_is_weak_barbed_simulation[where TRel="TRel"]
            indRelTEQ_impl_TRel_is_weak_barbed_contrasimulation[where TRel="TRel"]
    by blast

text ‹If indRelRTPO, indRelLTPO, or indRelTEQ is a correspondence simulation then so is the
        corresponding target term relation.›

lemma (in encoding) rel_with_target_impl_transC_TRel_is_weak_reduction_correspondence_simulation:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes corSim: "weak_reduction_correspondence_simulation Rel (STCal Source Target)"
      and target: "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
      and trel:   "T1 T2. (TargetTerm T1, TargetTerm T2)  Rel  (T1, T2)  TRel+"
  shows "weak_reduction_correspondence_simulation (TRel+) Target"
proof -
  from corSim target trel have A: "weak_reduction_simulation (TRel+) Target"
      using rel_with_target_impl_transC_TRel_is_weak_reduction_simulation[where TRel="TRel"
             and Rel="Rel"]
    by blast
  moreover have "P Q Q'. (P, Q)  TRel+  Q Target* Q'
                  (P'' Q''. P Target* P''  Q' Target* Q''  (P'', Q'')  TRel+)"
  proof clarify
    fix TP TQ TQ'
    assume "(TP, TQ)  TRel+" and "TQ Target* TQ'"
    thus "TP'' TQ''. TP Target* TP''  TQ' Target* TQ''  (TP'', TQ'')  TRel+"
    proof (induct arbitrary: TQ')
      fix TQ TQ'
      assume "(TP, TQ)  TRel"
      with target have "(TargetTerm TP, TargetTerm TQ)  Rel"
        by blast
      moreover assume "TQ Target* TQ'"
      hence "TargetTerm TQ (STCal Source Target)* (TargetTerm TQ')"
        by (simp add: STCal_steps)
      ultimately obtain P'' Q'' where A2: "TargetTerm TP (STCal Source Target)* P''"
        and A3: "TargetTerm TQ' (STCal Source Target)* Q''" and A4: "(P'', Q'')  Rel"
          using corSim
        by blast
      from A2 obtain TP'' where A5: "TP Target* TP''" and A6: "TP'' ∈T P''"
        by (auto simp add: STCal_steps)
      from A3 obtain TQ'' where A7: "TQ' Target* TQ''" and A8: "TQ'' ∈T Q''"
        by (auto simp add: STCal_steps)
      from A4 A6 A8 trel have "(TP'', TQ'')  TRel+"
        by blast
      with A5 A7
      show "TP'' TQ''. TP Target* TP''  TQ' Target* TQ''  (TP'', TQ'')  TRel+"
        by blast
    next
      case (step TQ TR TR')
      assume "TQ'. TQ Target* TQ' TP'' TQ''. TP Target* TP''  TQ' Target* TQ''
               (TP'', TQ'')  TRel+"
      moreover assume "(TQ, TR)  TRel"
      hence "TR'. TR Target* TR'
              (TQ'' TR''. TQ Target* TQ''  TR' Target* TR''  (TQ'', TR'')  TRel+)"
      proof clarify
        fix TR'
        assume "(TQ, TR)  TRel"
        with target have "(TargetTerm TQ, TargetTerm TR)  Rel"
          by simp
        moreover assume "TR Target* TR'"
        hence "TargetTerm TR (STCal Source Target)* (TargetTerm TR')"
          by (simp add: STCal_steps)
        ultimately obtain Q'' R'' where B1: "TargetTerm TQ (STCal Source Target)* Q''"
         and B2: "TargetTerm TR' (STCal Source Target)* R''" and B3: "(Q'', R'')  Rel"
            using corSim
          by blast
        from B1 obtain TQ'' where B4: "TQ'' ∈T Q''" and B5: "TQ Target* TQ''"
          by (auto simp add: STCal_steps)
        from B2 obtain TR'' where B6: "TR'' ∈T R''" and B7: "TR' Target* TR''"
          by (auto simp add: STCal_steps)
        from B3 B4 B6 trel have "(TQ'', TR'')  TRel+"
          by simp
        with B5 B7
        show "TQ'' TR''. TQ Target* TQ''  TR' Target* TR''  (TQ'', TR'')  TRel+"
          by blast
      qed
      moreover have "trans (TRel+)"
        by simp
      moreover assume "TR Target* TR'"
      ultimately
      show "TP'' TR''. TP Target* TP''  TR' Target* TR''  (TP'', TR'')  TRel+"
        using A reduction_correspondence_simulation_condition_trans[where Rel="TRel+"
                and Cal="Target"]
        by blast
    qed
  qed
  ultimately show ?thesis
    by simp
qed

lemma (in encoding) indRelRTPO_impl_TRel_is_weak_reduction_correspondence_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes cSim: "weak_reduction_correspondence_simulation (indRelRTPO TRel) (STCal Source Target)"
  shows "weak_reduction_correspondence_simulation (TRel+) Target"
      using cSim indRelRTPO.target[where TRel="TRel"] indRelRTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_weak_reduction_correspondence_simulation[where
              Rel="indRelRTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding) indRelLTPO_impl_TRel_is_weak_reduction_correspondence_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes cSim: "weak_reduction_correspondence_simulation (indRelLTPO TRel) (STCal Source Target)"
  shows "weak_reduction_correspondence_simulation (TRel+) Target"
      using cSim indRelLTPO.target[where TRel="TRel"] indRelLTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_weak_reduction_correspondence_simulation[where
              Rel="indRelLTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding)
  rel_with_target_impl_reflC_transC_TRel_is_weak_reduction_correspondence_simulation:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes corSim: "weak_reduction_correspondence_simulation Rel (STCal Source Target)"
      and target: "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
      and trel:   "T1 T2. (TargetTerm T1, TargetTerm T2)  Rel  (T1, T2)  TRel*"
  shows "weak_reduction_correspondence_simulation (TRel*) Target"
proof -
  from corSim target trel have A: "weak_reduction_simulation (TRel*) Target"
      using rel_with_target_impl_reflC_transC_TRel_is_weak_reduction_simulation[where TRel="TRel"
             and Rel="Rel"]
    by blast
  moreover have "P Q Q'. (P, Q)  TRel*  Q Target* Q' 
    (P'' Q''. P Target* P''  Q' Target* Q''  (P'', Q'')  TRel*)"
  proof clarify
    fix TP TQ TQ'
    assume "(TP, TQ)  TRel*" and "TQ Target* TQ'"
    thus "TP'' TQ''. TP Target* TP''  TQ' Target* TQ''  (TP'', TQ'')  TRel*"
    proof (induct arbitrary: TQ')
      fix TQ'
      assume "TP Target* TQ'"
      moreover have "TQ' Target* TQ'"
        by (simp add: steps_refl)
      moreover have "(TQ', TQ')  TRel*"
        by simp
      ultimately show "TP'' TQ''. TP Target* TP''  TQ' Target* TQ''  (TP'', TQ'')  TRel*"
        by blast
    next
      case (step TQ TR TR')
      assume "TQ'. TQ Target* TQ' TP'' TQ''. TP Target* TP''  TQ' Target* TQ''
               (TP'', TQ'')  TRel*"
      moreover assume "(TQ, TR)  TRel"
      with corSim have "TR'. TR Target* TR'  TQ'' TR''. TQ Target* TQ''
                         TR' Target* TR''  (TQ'', TR'')  TRel*"
      proof clarify
        fix TR'
        assume "(TQ, TR)  TRel"
        with target have "(TargetTerm TQ, TargetTerm TR)  Rel"
          by simp
        moreover assume "TR Target* TR'"
        hence "TargetTerm TR (STCal Source Target)* (TargetTerm TR')"
          by (simp add: STCal_steps)
        ultimately obtain Q'' R'' where B1: "TargetTerm TQ (STCal Source Target)* Q''"
         and B2: "TargetTerm TR' (STCal Source Target)* R''" and B3: "(Q'', R'')  Rel"
            using corSim
          by blast
        from B1 obtain TQ'' where B4: "TQ'' ∈T Q''" and B5: "TQ Target* TQ''"
          by (auto simp add: STCal_steps)
        from B2 obtain TR'' where B6: "TR'' ∈T R''" and B7: "TR' Target* TR''"
          by (auto simp add: STCal_steps)
        from B3 B4 B6 trel have "(TQ'', TR'')  TRel*"
          by simp
        with B5 B7
        show "TQ'' TR''. TQ Target* TQ''  TR' Target* TR''  (TQ'', TR'')  TRel*"
          by blast
      qed
      moreover assume "TR Target* TR'"
      moreover have "trans (TRel*)"
          using trans_rtrancl[of TRel]
        by simp
      ultimately show "TP'' TR''. TP Target* TP''  TR' Target* TR''  (TP'', TR'')  TRel*"
        using A reduction_correspondence_simulation_condition_trans[where Rel="TRel*"
                and Cal="Target"]
        by blast
    qed
  qed
  ultimately show ?thesis
    by simp
qed

lemma (in encoding) indRelTEQ_impl_TRel_is_weak_reduction_correspondence_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes corSim: "weak_reduction_correspondence_simulation (indRelTEQ TRel) (STCal Source Target)"
  shows "weak_reduction_correspondence_simulation (TRel*) Target"
      using corSim indRelTEQ.target[where TRel="TRel"] indRelTEQ_to_TRel(4)[where TRel="TRel"]
            trans_closure_of_TRel_refl_cond
            rel_with_target_impl_reflC_transC_TRel_is_weak_reduction_correspondence_simulation[
              where Rel="indRelTEQ TRel" and TRel="TRel"]
    by blast

lemma (in encoding_wrt_barbs) indRelRTPO_impl_TRel_is_weak_barbed_correspondence_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes corSim: "weak_barbed_correspondence_simulation (indRelRTPO TRel) (STCalWB SWB TWB)"
  shows "weak_barbed_correspondence_simulation (TRel+) TWB"
proof
  from corSim show "weak_reduction_correspondence_simulation (TRel+) (Calculus TWB)"
      using indRelRTPO_impl_TRel_is_weak_reduction_correspondence_simulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from corSim show "rel_weakly_respects_barbs (TRel+) TWB"
      using indRelRTPO_impl_TRel_weakly_respects_barbs[where TRel="TRel"]
            weak_respection_of_barbs_and_closures(3)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelLTPO_impl_TRel_is_weak_barbed_correspondence_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes corSim: "weak_barbed_correspondence_simulation (indRelLTPO TRel) (STCalWB SWB TWB)"
  shows "weak_barbed_correspondence_simulation (TRel+) TWB"
proof
  from corSim show "weak_reduction_correspondence_simulation (TRel+) (Calculus TWB)"
      using indRelLTPO_impl_TRel_is_weak_reduction_correspondence_simulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from corSim show "rel_weakly_respects_barbs (TRel+) TWB"
      using indRelLTPO_impl_TRel_weakly_respects_barbs[where TRel="TRel"]
            weak_respection_of_barbs_and_closures(3)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelTEQ_impl_TRel_is_weak_barbed_correspondence_simulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes corSim: "weak_barbed_correspondence_simulation (indRelTEQ TRel) (STCalWB SWB TWB)"
  shows "weak_barbed_correspondence_simulation (TRel*) TWB"
proof
  from corSim show "weak_reduction_correspondence_simulation (TRel*) (Calculus TWB)"
      using indRelTEQ_impl_TRel_is_weak_reduction_correspondence_simulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from corSim show "rel_weakly_respects_barbs (TRel*) TWB"
      using indRelTEQ_impl_TRel_weakly_respects_barbs[where TRel="TRel"]
            weak_respection_of_barbs_and_closures(5)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

text ‹If indRelRTPO, indRelLTPO, or indRelTEQ is a bisimulation then so is the corresponding
        target term relation.›

lemma (in encoding) rel_with_target_impl_transC_TRel_is_weak_reduction_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes bisim:  "weak_reduction_bisimulation Rel (STCal Source Target)"
      and target: "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
      and trel:   "T1 T2. (TargetTerm T1, TargetTerm T2)  Rel  (T1, T2)  TRel+"
  shows "weak_reduction_bisimulation (TRel+) Target"
proof
  from bisim target trel show "weak_reduction_simulation (TRel+) Target"
      using rel_with_target_impl_transC_TRel_is_weak_reduction_simulation[where TRel="TRel"
             and Rel="Rel"]
    by blast
next
  show "P Q Q'. (P, Q)  TRel+  Q Target* Q'  (P'. P Target* P'  (P', Q')  TRel+)"
  proof clarify
    fix TP TQ TQ'
    assume "(TP, TQ)  TRel+" and "TQ Target* TQ'"
    thus "TP'. TP Target* TP'  (TP', TQ')  TRel+"
    proof (induct arbitrary: TQ')
      fix TQ TQ'
      assume "(TP, TQ)  TRel"
      with target have "(TargetTerm TP, TargetTerm TQ)  Rel"
        by simp
      moreover assume "TQ Target* TQ'"
      hence "TargetTerm TQ (STCal Source Target)* (TargetTerm TQ')"
        by (simp add: STCal_steps)
      ultimately obtain P' where A2: "TargetTerm TP (STCal Source Target)* P'"
                             and A3: "(P', TargetTerm TQ')  Rel"
          using bisim
        by blast
      from A2 obtain TP' where A4: "TP Target* TP'" and A5: "TP' ∈T P'"
        by (auto simp add: STCal_steps)
      from A3 A5 trel have "(TP', TQ')  TRel+"
        by simp
      with A4 show "TP'. TP Target* TP'  (TP', TQ')  TRel+"
        by blast
    next
      case (step TQ TR TR')
      assume "(TQ, TR)  TRel"
      with target have "(TargetTerm TQ, TargetTerm TR)  Rel"
        by simp
      moreover assume "TR Target* TR'"
      hence "TargetTerm TR (STCal Source Target)* (TargetTerm TR')"
        by (simp add: STCal_steps)
      ultimately obtain Q' where B1: "TargetTerm TQ (STCal Source Target)* Q'"
                             and B2: "(Q', TargetTerm TR')  Rel"
          using bisim
        by blast
      from B1 obtain TQ' where B3: "TQ' ∈T Q'" and B4: "TQ Target* TQ'"
        by (auto simp add: STCal_steps)
      assume "TQ'. TQ Target* TQ'  TP'. TP Target* TP'  (TP', TQ')  TRel+"
      with B4 obtain TP' where B5: "TP Target* TP'" and B6: "(TP', TQ')  TRel+"
        by blast
      from B2 B3 trel have "(TQ', TR')  TRel+"
        by simp
      with B6 have "(TP', TR')  TRel+"
        by simp
      with B5 show "TP'. TP Target* TP'  (TP', TR')  TRel+"
        by blast
    qed
  qed
qed

lemma (in encoding) indRelRTPO_impl_TRel_is_weak_reduction_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes bisim: "weak_reduction_bisimulation (indRelRTPO TRel) (STCal Source Target)"
  shows "weak_reduction_bisimulation (TRel+) Target"
      using bisim indRelRTPO.target[where TRel="TRel"] indRelRTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_weak_reduction_bisimulation[where
             Rel="indRelRTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding) indRelLTPO_impl_TRel_is_weak_reduction_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes bisim: "weak_reduction_bisimulation (indRelLTPO TRel) (STCal Source Target)"
  shows "weak_reduction_bisimulation (TRel+) Target"
      using bisim indRelLTPO.target[where TRel="TRel"] indRelLTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_weak_reduction_bisimulation[where
             Rel="indRelLTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding) rel_with_target_impl_reflC_transC_TRel_is_weak_reduction_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes bisim:  "weak_reduction_bisimulation Rel (STCal Source Target)"
      and target: "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
      and trel:   "T1 T2. (TargetTerm T1, TargetTerm T2)  Rel  (T1, T2)  TRel*"
  shows "weak_reduction_bisimulation (TRel*) Target"
proof
  from bisim target trel show "weak_reduction_simulation (TRel*) Target"
      using rel_with_target_impl_reflC_transC_TRel_is_weak_reduction_simulation[where TRel="TRel"
             and Rel="Rel"]
    by blast
next
  show "P Q Q'. (P, Q)  TRel*  Q Target* Q'  (P'. P Target* P'  (P', Q')  TRel*)"
  proof clarify
    fix TP TQ TQ'
    assume "(TP, TQ)  TRel*" and "TQ Target* TQ'"
    thus "TP'. TP Target* TP'  (TP', TQ')  TRel*"
    proof (induct arbitrary: TQ')
      fix TQ'
      assume "TP Target* TQ'"
      moreover have "(TQ', TQ')  TRel*"
        by simp
      ultimately show "TP'. TP Target* TP'  (TP', TQ')  TRel*"
        by blast
    next
      case (step TQ TR TR')
      assume "(TQ, TR)  TRel"
      with target have "(TargetTerm TQ, TargetTerm TR)  Rel"
        by simp
      moreover assume "TR Target* TR'"
      hence "TargetTerm TR (STCal Source Target)* (TargetTerm TR')"
        by (simp add: STCal_steps)
      ultimately obtain Q' where B1: "TargetTerm TQ (STCal Source Target)* Q'"
                             and B2: "(Q', TargetTerm TR')  Rel"
          using bisim
        by blast
      from B1 obtain TQ' where B3: "TQ' ∈T Q'" and B4: "TQ Target* TQ'"
        by (auto simp add: STCal_steps)
      assume "TQ'. TQ Target* TQ'  TP'. TP Target* TP'  (TP', TQ')  TRel*"
      with B4 obtain TP' where B5: "TP Target* TP'" and B6: "(TP', TQ')  TRel*"
        by blast
      from B2 B3 trel have "(TQ', TR')  TRel*"
        by simp
      with B6 have "(TP', TR')  TRel*"
        by simp
      with B5 show "TP'. TP Target* TP'  (TP', TR')  TRel*"
        by blast
    qed
  qed
qed

lemma (in encoding) indRelTEQ_impl_TRel_is_weak_reduction_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes bisim: "weak_reduction_bisimulation (indRelTEQ TRel) (STCal Source Target)"
  shows "weak_reduction_bisimulation (TRel*) Target"
      using bisim indRelTEQ.target[where TRel="TRel"] indRelTEQ_to_TRel(4)[where TRel="TRel"]
            trans_closure_of_TRel_refl_cond
            rel_with_target_impl_reflC_transC_TRel_is_weak_reduction_bisimulation[where
             Rel="indRelTEQ TRel" and TRel="TRel"]
    by blast

lemma (in encoding) rel_with_target_impl_transC_TRel_is_strong_reduction_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes bisim:  "strong_reduction_bisimulation Rel (STCal Source Target)"
      and target: "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
      and trel:   "T1 T2. (TargetTerm T1, TargetTerm T2)  Rel  (T1, T2)  TRel+"
  shows "strong_reduction_bisimulation (TRel+) Target"
proof
  from bisim target trel show "strong_reduction_simulation (TRel+) Target"
      using rel_with_target_impl_transC_TRel_is_strong_reduction_simulation[where Rel="Rel"
              and TRel="TRel"]
    by blast
next
  show "P Q Q'. (P, Q)  TRel+  Q Target Q'  (P'. P Target P'  (P', Q')  TRel+)"
  proof clarify
    fix TP TQ TQ'
    assume "(TP, TQ)  TRel+" and "TQ Target TQ'"
    thus "TP'. TP Target TP'  (TP', TQ')  TRel+"
    proof (induct arbitrary: TQ')
      fix TQ TQ'
      assume "(TP, TQ)  TRel"
      with target have "(TargetTerm TP, TargetTerm TQ)  Rel"
        by simp
      moreover assume "TQ Target TQ'"
      hence "TargetTerm TQ (STCal Source Target) (TargetTerm TQ')"
        by (simp add: STCal_step)
      ultimately obtain P' where A2: "TargetTerm TP (STCal Source Target) P'"
                             and A3: "(P', TargetTerm TQ')  Rel"
          using bisim
        by blast
      from A2 obtain TP' where A4: "TP Target TP'" and A5: "TP' ∈T P'"
        by (auto simp add: STCal_step)
      from A3 A5 trel have "(TP', TQ')  TRel+"
        by simp
      with A4 show "TP'. TP Target TP'  (TP', TQ')  TRel+"
        by blast
    next
      case (step TQ TR TR')
      assume "(TQ, TR)  TRel"
      with target have "(TargetTerm TQ, TargetTerm TR)  Rel"
        by simp
      moreover assume "TR Target TR'"
      hence "TargetTerm TR (STCal Source Target) (TargetTerm TR')"
        by (simp add: STCal_step)
      ultimately obtain Q' where B1: "TargetTerm TQ (STCal Source Target) Q'"
                             and B2: "(Q', TargetTerm TR')  Rel"
          using bisim
        by blast
      from B1 obtain TQ' where B3: "TQ' ∈T Q'" and B4: "TQ Target TQ'"
        by (auto simp add: STCal_step)
      assume "TQ'. TQ Target TQ'  TP'. TP Target TP'  (TP', TQ')  TRel+"
      with B4 obtain TP' where B5: "TP Target TP'" and B6: "(TP', TQ')  TRel+"
        by blast
      from B2 B3 trel have "(TQ', TR')  TRel+"
        by simp
      with B6 have "(TP', TR')  TRel+"
        by simp
      with B5 show "TP'. TP Target TP'  (TP', TR')  TRel+"
        by blast
    qed
  qed
qed

lemma (in encoding) indRelRTPO_impl_TRel_is_strong_reduction_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes bisim: "strong_reduction_bisimulation (indRelRTPO TRel) (STCal Source Target)"
  shows "strong_reduction_bisimulation (TRel+) Target"
      using bisim indRelRTPO.target[where TRel="TRel"] indRelRTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_strong_reduction_bisimulation[where
             Rel="indRelRTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding) indRelLTPO_impl_TRel_is_strong_reduction_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes bisim: "strong_reduction_bisimulation (indRelLTPO TRel) (STCal Source Target)"
  shows "strong_reduction_bisimulation (TRel+) Target"
      using bisim indRelLTPO.target[where TRel="TRel"] indRelLTPO_to_TRel(4)[where TRel="TRel"]
            rel_with_target_impl_transC_TRel_is_strong_reduction_bisimulation[where
             Rel="indRelLTPO TRel" and TRel="TRel"]
    by blast

lemma (in encoding) rel_with_target_impl_reflC_transC_TRel_is_strong_reduction_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes bisim:  "strong_reduction_bisimulation Rel (STCal Source Target)"
      and target: "T1 T2. (T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  Rel"
      and trel:   "T1 T2. (TargetTerm T1, TargetTerm T2)  Rel  (T1, T2)  TRel*"
  shows "strong_reduction_bisimulation (TRel*) Target"
proof
  from bisim target trel show "strong_reduction_simulation (TRel*) Target"
      using rel_with_target_impl_reflC_transC_TRel_is_strong_reduction_simulation[where Rel="Rel"
              and TRel="TRel"]
    by blast
next
  show "P Q Q'. (P, Q)  TRel*  Q Target Q'  (P'. P Target P'  (P', Q')  TRel*)"
  proof clarify
    fix TP TQ TQ'
    assume "(TP, TQ)  TRel*" and "TQ Target TQ'"
    thus "TP'. TP Target TP'  (TP', TQ')  TRel*"
    proof (induct arbitrary: TQ')
      fix TQ'
      assume "TP Target TQ'"
      thus "TP'. TP Target TP'  (TP', TQ')  TRel*"
        by blast
    next
      case (step TQ TR TR')
      assume "(TQ, TR)  TRel"
      with target have "(TargetTerm TQ, TargetTerm TR)  Rel"
        by simp
      moreover assume "TR Target TR'"
      hence "TargetTerm TR (STCal Source Target) (TargetTerm TR')"
        by (simp add: STCal_step)
      ultimately obtain Q' where B1: "TargetTerm TQ (STCal Source Target) Q'"
                             and B2: "(Q', TargetTerm TR')  Rel"
          using bisim
        by blast
      from B1 obtain TQ' where B3: "TQ' ∈T Q'" and B4: "TQ Target TQ'"
        by (auto simp add: STCal_step)
      assume "TQ'. TQ Target TQ'  TP'. TP Target TP'  (TP', TQ')  TRel*"
      with B4 obtain TP' where B5: "TP Target TP'" and B6: "(TP', TQ')  TRel*"
        by blast
      from B2 B3 trel have "(TQ', TR')  TRel*"
        by simp
      with B6 have "(TP', TR')  TRel*"
        by simp
      with B5 show "TP'. TP Target TP'  (TP', TR')  TRel*"
        by blast
    qed
  qed
qed

lemma (in encoding) indRelTEQ_impl_TRel_is_strong_reduction_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes bisim: "strong_reduction_bisimulation (indRelTEQ TRel) (STCal Source Target)"
  shows "strong_reduction_bisimulation (TRel*) Target"
      using bisim indRelTEQ.target[where TRel="TRel"] indRelTEQ_to_TRel(4)[where TRel="TRel"]
            trans_closure_of_TRel_refl_cond
            rel_with_target_impl_reflC_transC_TRel_is_strong_reduction_bisimulation[where
             Rel="indRelTEQ TRel" and TRel="TRel"]
    by blast

lemma (in encoding_wrt_barbs) indRelRTPO_impl_TRel_is_weak_barbed_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes bisim: "weak_barbed_bisimulation (indRelRTPO TRel) (STCalWB SWB TWB)"
  shows "weak_barbed_bisimulation (TRel+) TWB"
proof
  from bisim show "weak_reduction_bisimulation (TRel+) (Calculus TWB)"
      using indRelRTPO_impl_TRel_is_weak_reduction_bisimulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from bisim show "rel_weakly_respects_barbs (TRel+) TWB"
      using indRelRTPO_impl_TRel_weakly_respects_barbs[where TRel="TRel"]
            weak_respection_of_barbs_and_closures(3)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelLTPO_impl_TRel_is_weak_barbed_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes bisim: "weak_barbed_bisimulation (indRelLTPO TRel) (STCalWB SWB TWB)"
  shows "weak_barbed_bisimulation (TRel+) TWB"
proof
  from bisim show "weak_reduction_bisimulation (TRel+) (Calculus TWB)"
      using indRelLTPO_impl_TRel_is_weak_reduction_bisimulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from bisim show "rel_weakly_respects_barbs (TRel+) TWB"
      using indRelLTPO_impl_TRel_weakly_respects_barbs[where TRel="TRel"]
            weak_respection_of_barbs_and_closures(3)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelTEQ_impl_TRel_is_weak_barbed_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes bisim: "weak_barbed_bisimulation (indRelTEQ TRel) (STCalWB SWB TWB)"
  shows "weak_barbed_bisimulation (TRel*) TWB"
proof
  from bisim show "weak_reduction_bisimulation (TRel*) (Calculus TWB)"
      using indRelTEQ_impl_TRel_is_weak_reduction_bisimulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from bisim show "rel_weakly_respects_barbs (TRel*) TWB"
      using indRelTEQ_impl_TRel_weakly_respects_barbs[where TRel="TRel"]
            weak_respection_of_barbs_and_closures(5)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelRTPO_impl_TRel_is_strong_barbed_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes bisim: "strong_barbed_bisimulation (indRelRTPO TRel) (STCalWB SWB TWB)"
  shows "strong_barbed_bisimulation (TRel+) TWB"
proof
  from bisim show "strong_reduction_bisimulation (TRel+) (Calculus TWB)"
      using indRelRTPO_impl_TRel_is_strong_reduction_bisimulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from bisim show "rel_respects_barbs (TRel+) TWB"
      using indRelRTPO_impl_TRel_respects_barbs[where TRel="TRel"]
            respection_of_barbs_and_closures(3)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelLTPO_impl_TRel_is_strong_barbed_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes bisim: "strong_barbed_bisimulation (indRelLTPO TRel) (STCalWB SWB TWB)"
  shows "strong_barbed_bisimulation (TRel+) TWB"
proof
  from bisim refl show "strong_reduction_bisimulation (TRel+) (Calculus TWB)"
      using indRelLTPO_impl_TRel_is_strong_reduction_bisimulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from bisim show "rel_respects_barbs (TRel+) TWB"
      using indRelLTPO_impl_TRel_respects_barbs[where TRel="TRel"]
            respection_of_barbs_and_closures(3)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelTEQ_impl_TRel_is_strong_barbed_bisimulation:
  fixes TRel :: "('procT × 'procT) set"
  assumes bisim: "strong_barbed_bisimulation (indRelTEQ TRel) (STCalWB SWB TWB)"
  shows "strong_barbed_bisimulation (TRel*) TWB"
proof
  from bisim refl show "strong_reduction_bisimulation (TRel*) (Calculus TWB)"
      using indRelTEQ_impl_TRel_is_strong_reduction_bisimulation[where TRel="TRel"]
    by (simp add: STCalWB_def calS calT)
next
  from bisim show "rel_respects_barbs (TRel*) TWB"
      using indRelTEQ_impl_TRel_respects_barbs[where TRel="TRel"]
            respection_of_barbs_and_closures(5)[where Rel="TRel" and CWB="TWB"]
    by blast
qed

subsection ‹Relations Induced by the Encoding and Relations on Source Terms and Target Terms›

text ‹Some encodability like e.g. full abstraction are defined w.r.t. a relation on source terms
        and a relation on target terms. To analyse such criteria we include these two relations in
        the considered relation on the disjoint union of source and target terms.›

inductive_set (in encoding) indRelRST
    :: "('procS × 'procS) set  ('procT × 'procT) set
         ((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
    for SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  where
  encR:   "(SourceTerm S, TargetTerm (S))  indRelRST SRel TRel" |
  source: "(S1, S2)  SRel  (SourceTerm S1, SourceTerm S2)  indRelRST SRel TRel" |
  target: "(T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  indRelRST SRel TRel"

abbreviation (in encoding) indRelRSTinfix
    :: "('procS, 'procT) Proc  ('procS × 'procS) set  ('procT × 'procT) set
         ('procS, 'procT) Proc  bool" ("_ ℛ⟦⋅⟧R<_,_> _" [75, 75, 75, 75] 80)
  where
  "P ℛ⟦⋅⟧R<SRel,TRel> Q  (P, Q)  indRelRST SRel TRel"

inductive_set (in encoding) indRelRSTPO
    :: "('procS × 'procS) set  ('procT × 'procT) set
         ((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
    for SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  where
  encR:   "(SourceTerm S, TargetTerm (S))  indRelRSTPO SRel TRel" |
  source: "(S1, S2)  SRel  (SourceTerm S1, SourceTerm S2)  indRelRSTPO SRel TRel" |
  target: "(T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  indRelRSTPO SRel TRel" |
  trans:  "(P, Q)  indRelRSTPO SRel TRel; (Q, R)  indRelRSTPO SRel TRel
            (P, R)  indRelRSTPO SRel TRel"

abbreviation (in encoding) indRelRSTPOinfix ::
    "('procS, 'procT) Proc  ('procS × 'procS) set  ('procT × 'procT) set
      ('procS, 'procT) Proc  bool" ("_ ≲⟦⋅⟧R<_,_> _" [75, 75, 75, 75] 80)
  where
  "P ≲⟦⋅⟧R<SRel,TRel> Q  (P, Q)  indRelRSTPO SRel TRel"

lemma (in encoding) indRelRSTPO_refl:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes reflS: "refl SRel"
      and reflT: "refl TRel"
  shows "refl (indRelRSTPO SRel TRel)"
    unfolding refl_on_def
proof auto
  fix P
  show "P ≲⟦⋅⟧R<SRel,TRel> P"
  proof (cases P)
    case (SourceTerm SP)
    assume "SP ∈S P"
    with reflS show "P ≲⟦⋅⟧R<SRel,TRel> P"
        unfolding refl_on_def
      by (simp add: indRelRSTPO.source)
  next
    case (TargetTerm TP)
    assume "TP ∈T P"
    with reflT show "P ≲⟦⋅⟧R<SRel,TRel> P"
        unfolding refl_on_def
      by (simp add: indRelRSTPO.target)
  qed
qed

lemma (in encoding) indRelRSTPO_trans:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  shows "trans (indRelRSTPO SRel TRel)"
    unfolding trans_def
proof clarify
  fix P Q R
  assume "P ≲⟦⋅⟧R<SRel,TRel> Q" and "Q ≲⟦⋅⟧R<SRel,TRel> R"
  thus "P ≲⟦⋅⟧R<SRel,TRel> R"
    by (rule indRelRSTPO.trans)
qed

lemma (in encoding) refl_trans_closure_of_indRelRST:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes reflS: "refl SRel"
      and reflT: "refl TRel"
  shows "indRelRSTPO SRel TRel = (indRelRST SRel TRel)*"
proof auto
  fix P Q
  assume "P ≲⟦⋅⟧R<SRel,TRel> Q"
  thus "(P, Q)  (indRelRST SRel TRel)*"
  proof induct
    case (encR S)
    show "(SourceTerm S, TargetTerm (S))  (indRelRST SRel TRel)*"
        using indRelRST.encR[of S SRel TRel]
      by simp
  next
    case (source S1 S2)
    assume "(S1, S2)  SRel"
    thus "(SourceTerm S1, SourceTerm S2)  (indRelRST SRel TRel)*"
        using indRelRST.source[of S1 S2 SRel TRel]
      by simp
  next
    case (target T1 T2)
    assume "(T1, T2)  TRel"
    thus "(TargetTerm T1, TargetTerm T2)  (indRelRST SRel TRel)*"
        using indRelRST.target[of T1 T2 TRel SRel]
      by simp
  next
    case (trans P Q R)
    assume "(P, Q)  (indRelRST SRel TRel)*" and "(Q, R)  (indRelRST SRel TRel)*"
    thus "(P, R)  (indRelRST SRel TRel)*"
      by simp
  qed
next
  fix P Q
  assume "(P, Q)  (indRelRST SRel TRel)*"
  thus "P ≲⟦⋅⟧R<SRel,TRel> Q"
  proof induct
    from reflS reflT show "P ≲⟦⋅⟧R<SRel,TRel> P"
        using indRelRSTPO_refl[of SRel TRel]
        unfolding refl_on_def
      by simp
  next
    case (step Q R)
    assume "P ≲⟦⋅⟧R<SRel,TRel> Q"
    moreover assume "Q ℛ⟦⋅⟧R<SRel,TRel> R"
    hence "Q ≲⟦⋅⟧R<SRel,TRel> R"
      by (induct, simp_all add: indRelRSTPO.intros)
    ultimately show "P ≲⟦⋅⟧R<SRel,TRel> R"
      by (rule indRelRSTPO.trans)
  qed
qed

inductive_set (in encoding) indRelLST
    :: "('procS × 'procS) set  ('procT × 'procT) set
         ((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
    for SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  where
  encL:   "(TargetTerm (S), SourceTerm S)  indRelLST SRel TRel" |
  source: "(S1, S2)  SRel  (SourceTerm S1, SourceTerm S2)  indRelLST SRel TRel" |
  target: "(T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  indRelLST SRel TRel"

abbreviation (in encoding) indRelLSTinfix
    :: "('procS, 'procT) Proc  ('procS × 'procS) set  ('procT × 'procT) set
         ('procS, 'procT) Proc  bool" ("_ ℛ⟦⋅⟧L<_,_> _" [75, 75, 75, 75] 80)
  where
  "P ℛ⟦⋅⟧L<SRel,TRel> Q  (P, Q)  indRelLST SRel TRel"

inductive_set (in encoding) indRelLSTPO
    :: "('procS × 'procS) set  ('procT × 'procT) set
         ((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
    for SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  where
  encL:   "(TargetTerm (S), SourceTerm S)  indRelLSTPO SRel TRel" |
  source: "(S1, S2)  SRel  (SourceTerm S1, SourceTerm S2)  indRelLSTPO SRel TRel" |
  target: "(T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  indRelLSTPO SRel TRel" |
  trans:  "(P, Q)  indRelLSTPO SRel TRel; (Q, R)  indRelLSTPO SRel TRel
            (P, R)  indRelLSTPO SRel TRel"

abbreviation (in encoding) indRelLSTPOinfix
    :: "('procS, 'procT) Proc  ('procS × 'procS) set  ('procT × 'procT) set
         ('procS, 'procT) Proc  bool" ("_ ≲⟦⋅⟧L<_,_> _" [75, 75, 75, 75] 80)
  where
  "P ≲⟦⋅⟧L<SRel,TRel> Q  (P, Q)  indRelLSTPO SRel TRel"

lemma (in encoding) indRelLSTPO_refl:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes reflS: "refl SRel"
      and reflT: "refl TRel"
  shows "refl (indRelLSTPO SRel TRel)"
    unfolding refl_on_def
proof auto
  fix P
  show "P ≲⟦⋅⟧L<SRel,TRel> P"
  proof (cases P)
    case (SourceTerm SP)
    assume "SP ∈S P"
    with reflS show "P ≲⟦⋅⟧L<SRel,TRel> P"
        unfolding refl_on_def
      by (simp add: indRelLSTPO.source)
  next
    case (TargetTerm TP)
    assume "TP ∈T P"
    with reflT show "P ≲⟦⋅⟧L<SRel,TRel> P"
        unfolding refl_on_def
      by (simp add: indRelLSTPO.target)
  qed
qed

lemma (in encoding) indRelLSTPO_trans:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  shows "trans (indRelLSTPO SRel TRel)"
    unfolding trans_def
proof clarify
  fix P Q R
  assume "P ≲⟦⋅⟧L<SRel,TRel> Q" and "Q ≲⟦⋅⟧L<SRel,TRel> R"
  thus "P ≲⟦⋅⟧L<SRel,TRel> R"
    by (rule indRelLSTPO.trans)
qed

lemma (in encoding) refl_trans_closure_of_indRelLST:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes reflS: "refl SRel"
      and reflT: "refl TRel"
  shows "indRelLSTPO SRel TRel = (indRelLST SRel TRel)*"
proof auto
  fix P Q
  assume "P ≲⟦⋅⟧L<SRel,TRel> Q"
  thus "(P, Q)  (indRelLST SRel TRel)*"
  proof induct
    case (encL S)
    show "(TargetTerm (S), SourceTerm S)  (indRelLST SRel TRel)*"
        using indRelLST.encL[of S SRel TRel]
      by simp
  next
    case (source S1 S2)
    assume "(S1, S2)  SRel"
    thus "(SourceTerm S1, SourceTerm S2)  (indRelLST SRel TRel)*"
        using indRelLST.source[of S1 S2 SRel TRel]
      by simp
  next
    case (target T1 T2)
    assume "(T1, T2)  TRel"
    thus "(TargetTerm T1, TargetTerm T2)  (indRelLST SRel TRel)*"
        using indRelLST.target[of T1 T2 TRel SRel]
      by simp
  next
    case (trans P Q R)
    assume "(P, Q)  (indRelLST SRel TRel)*" and "(Q, R)  (indRelLST SRel TRel)*"
    thus "(P, R)  (indRelLST SRel TRel)*"
      by simp
  qed
next
  fix P Q
  assume "(P, Q)  (indRelLST SRel TRel)*"
  thus "P ≲⟦⋅⟧L<SRel,TRel> Q"
  proof induct
    from reflS reflT show " P ≲⟦⋅⟧L<SRel,TRel> P"
        using indRelLSTPO_refl[of SRel TRel]
        unfolding refl_on_def
      by simp
  next
    case (step Q R)
    assume "P ≲⟦⋅⟧L<SRel,TRel> Q"
    moreover assume "Q ℛ⟦⋅⟧L<SRel,TRel> R"
    hence "Q ≲⟦⋅⟧L<SRel,TRel> R"
      by (induct, simp_all add: indRelLSTPO.intros)
    ultimately show "P ≲⟦⋅⟧L<SRel,TRel> R"
      by (rule indRelLSTPO.trans)
  qed
qed

inductive_set (in encoding) indRelST
    :: "('procS × 'procS) set  ('procT × 'procT) set
         ((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
    for SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  where
  encR:   "(SourceTerm S, TargetTerm (S))  indRelST SRel TRel" |
  encL:   "(TargetTerm (S), SourceTerm S)  indRelST SRel TRel" |
  source: "(S1, S2)  SRel  (SourceTerm S1, SourceTerm S2)  indRelST SRel TRel" |
  target: "(T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  indRelST SRel TRel"

abbreviation (in encoding) indRelSTinfix
    :: "('procS, 'procT) Proc  ('procS × 'procS) set  ('procT × 'procT) set
         ('procS, 'procT) Proc  bool" ("_ ℛ⟦⋅⟧<_,_> _" [75, 75, 75, 75] 80)
  where
  "P ℛ⟦⋅⟧<SRel,TRel> Q  (P, Q)  indRelST SRel TRel"

lemma (in encoding) indRelST_symm:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes symmS: "sym SRel"
      and symmT: "sym TRel"
  shows "sym (indRelST SRel TRel)"
    unfolding sym_def
proof clarify
  fix P Q
  assume "(P, Q)  indRelST SRel TRel"
  thus "(Q, P)  indRelST SRel TRel"
  proof induct
    case (encR S)
    show "(TargetTerm (S), SourceTerm S)  indRelST SRel TRel"
      by (rule indRelST.encL)
  next
    case (encL S)
    show "(SourceTerm S, TargetTerm (S))  indRelST SRel TRel"
      by (rule indRelST.encR)
  next
    case (source S1 S2)
    assume "(S1, S2)  SRel"
    with symmS show "(SourceTerm S2, SourceTerm S1)  indRelST SRel TRel"
        unfolding sym_def
      by (simp add: indRelST.source)
  next
    case (target T1 T2)
    assume "(T1, T2)  TRel"
    with symmT show "(TargetTerm T2, TargetTerm T1)  indRelST SRel TRel"
        unfolding sym_def
      by (simp add: indRelST.target)
  qed
qed

inductive_set (in encoding) indRelSTEQ
    :: "('procS × 'procS) set  ('procT × 'procT) set
         ((('procS, 'procT) Proc) × (('procS, 'procT) Proc)) set"
    for SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  where
  encR:   "(SourceTerm S, TargetTerm (S))  indRelSTEQ SRel TRel" |
  encL:   "(TargetTerm (S), SourceTerm S)  indRelSTEQ SRel TRel" |
  source: "(S1, S2)  SRel  (SourceTerm S1, SourceTerm S2)  indRelSTEQ SRel TRel" |
  target: "(T1, T2)  TRel  (TargetTerm T1, TargetTerm T2)  indRelSTEQ SRel TRel" |
  trans:  "(P, Q)  indRelSTEQ SRel TRel; (Q, R)  indRelSTEQ SRel TRel
            (P, R)  indRelSTEQ SRel TRel"

abbreviation (in encoding) indRelSTEQinfix
    :: "('procS, 'procT) Proc  ('procS × 'procS) set  ('procT × 'procT) set
         ('procS, 'procT) Proc  bool" ("_ ∼⟦⋅⟧<_,_> _" [75, 75, 75, 75] 80)
  where
  "P ∼⟦⋅⟧<SRel,TRel> Q  (P, Q)  indRelSTEQ SRel TRel"

lemma (in encoding) indRelSTEQ_refl:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes reflT: "refl TRel"
  shows "refl (indRelSTEQ SRel TRel)"
    unfolding refl_on_def
proof auto
  fix P
  show "P ∼⟦⋅⟧<SRel,TRel> P"
  proof (cases P)
    case (SourceTerm SP)
    assume "SP ∈S P"
    moreover have "SourceTerm SP ∼⟦⋅⟧<SRel,TRel> TargetTerm (SP)"
      by (rule indRelSTEQ.encR)
    moreover have "TargetTerm (SP) ∼⟦⋅⟧<SRel,TRel> SourceTerm SP"
      by (rule indRelSTEQ.encL)
    ultimately show "P ∼⟦⋅⟧<SRel,TRel> P"
      by (simp add: indRelSTEQ.trans[where P="SourceTerm SP" and Q="TargetTerm (SP)"])
  next
    case (TargetTerm TP)
    assume "TP ∈T P"
    with reflT show "P ∼⟦⋅⟧<SRel,TRel> P"
        unfolding refl_on_def
      by (simp add: indRelSTEQ.target)
  qed
qed

lemma (in encoding) indRelSTEQ_symm:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes symmS: "sym SRel"
      and symmT: "sym TRel"
  shows "sym (indRelSTEQ SRel TRel)"
    unfolding sym_def
proof clarify
  fix P Q
  assume "P ∼⟦⋅⟧<SRel,TRel> Q"
  thus "Q ∼⟦⋅⟧<SRel,TRel> P"
  proof induct
    case (encR S)
    show "TargetTerm (S) ∼⟦⋅⟧<SRel,TRel> SourceTerm S"
      by (rule indRelSTEQ.encL)
  next
    case (encL S)
    show "SourceTerm S ∼⟦⋅⟧<SRel,TRel> TargetTerm (S)"
      by (rule indRelSTEQ.encR)
  next
    case (source S1 S2)
    assume "(S1, S2)  SRel"
    with symmS show "SourceTerm S2 ∼⟦⋅⟧<SRel,TRel> SourceTerm S1"
        unfolding sym_def
      by (simp add: indRelSTEQ.source)
  next
    case (target T1 T2)
    assume "(T1, T2)  TRel"
    with symmT show "TargetTerm T2 ∼⟦⋅⟧<SRel,TRel> TargetTerm T1"
        unfolding sym_def
      by (simp add: indRelSTEQ.target)
  next
    case (trans P Q R)
    assume "R ∼⟦⋅⟧<SRel,TRel> Q" and "Q ∼⟦⋅⟧<SRel,TRel> P"
    thus "R ∼⟦⋅⟧<SRel,TRel> P"
      by (rule indRelSTEQ.trans)
  qed
qed

lemma (in encoding) indRelSTEQ_trans:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  shows "trans (indRelSTEQ SRel TRel)"
    unfolding trans_def
proof clarify
  fix P Q R
  assume "P ∼⟦⋅⟧<SRel,TRel> Q" and "Q ∼⟦⋅⟧<SRel,TRel> R"
  thus "P ∼⟦⋅⟧<SRel,TRel> R"
    by (rule indRelSTEQ.trans)
qed

lemma (in encoding) refl_trans_closure_of_indRelST:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes reflT: "refl TRel"
  shows "indRelSTEQ SRel TRel = (indRelST SRel TRel)*"
proof auto
  fix P Q
  assume "P ∼⟦⋅⟧<SRel,TRel> Q"
  thus "(P, Q)  (indRelST SRel TRel)*"
  proof induct
    case (encR S)
    show "(SourceTerm S, TargetTerm (S))  (indRelST SRel TRel)*"
        using indRelST.encR[of S SRel TRel]
      by simp
  next
    case (encL S)
    show "(TargetTerm (S), SourceTerm S)  (indRelST SRel TRel)*"
        using indRelST.encL[of S SRel TRel]
      by simp
  next
    case (source S1 S2)
    assume "(S1, S2)  SRel"
    thus "(SourceTerm S1, SourceTerm S2)  (indRelST SRel TRel)*"
        using indRelST.source[of S1 S2 SRel TRel]
      by simp
  next
    case (target T1 T2)
    assume "(T1, T2)  TRel"
    thus "(TargetTerm T1, TargetTerm T2)  (indRelST SRel TRel)*"
        using indRelST.target[of T1 T2 TRel SRel]
      by simp
  next
    case (trans P Q R)
    assume "(P, Q)  (indRelST SRel TRel)*" and "(Q, R)  (indRelST SRel TRel)*"
    thus "(P, R)  (indRelST SRel TRel)*"
      by simp
  qed
next
  fix P Q
  assume "(P, Q)  (indRelST SRel TRel)*"
  thus "P ∼⟦⋅⟧<SRel,TRel> Q"
  proof induct
    from reflT show "P ∼⟦⋅⟧<SRel,TRel> P"
        using indRelSTEQ_refl[of TRel SRel]
        unfolding refl_on_def
      by simp
  next
    case (step Q R)
    assume "P ∼⟦⋅⟧<SRel,TRel> Q"
    moreover assume "Q ℛ⟦⋅⟧<SRel,TRel> R"
    hence "Q ∼⟦⋅⟧<SRel,TRel> R"
      by (induct, simp_all add: indRelSTEQ.intros)
    ultimately show "P ∼⟦⋅⟧<SRel,TRel> R"
      by (rule indRelSTEQ.trans)
  qed
qed

lemma (in encoding) refl_symm_trans_closure_of_indRelST:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes reflT: "refl TRel"
      and symmS: "sym SRel"
      and symmT: "sym TRel"
  shows "indRelSTEQ SRel TRel = (symcl ((indRelST SRel TRel)=))+"
proof -
  have "(symcl ((indRelST SRel TRel)=))+ = (symcl (indRelST SRel TRel))*"
    by (rule refl_symm_trans_closure_is_symm_refl_trans_closure[where Rel="indRelST SRel TRel"])
  moreover from symmS symmT have "symcl (indRelST SRel TRel) = indRelST SRel TRel"
      using indRelST_symm[where SRel="SRel" and TRel="TRel"]
            symm_closure_of_symm_rel[where Rel="indRelST SRel TRel"]
    by blast
  ultimately show "indRelSTEQ SRel TRel = (symcl ((indRelST SRel TRel)=))+"
      using reflT refl_trans_closure_of_indRelST[where SRel="SRel" and TRel="TRel"]
    by simp
qed

lemma (in encoding) symm_closure_of_indRelRST:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes reflT: "refl TRel"
      and symmS: "sym SRel"
      and symmT: "sym TRel"
  shows "indRelST SRel TRel = symcl (indRelRST SRel TRel)"
    and "indRelSTEQ SRel TRel = (symcl ((indRelRST SRel TRel)=))+"
proof -
  show "indRelST SRel TRel = symcl (indRelRST SRel TRel)"
  proof auto
    fix P Q
    assume "P ℛ⟦⋅⟧<SRel,TRel> Q"
    thus "(P, Q)  symcl (indRelRST SRel TRel)"
      by (induct, simp_all add: symcl_def indRelRST.intros)
  next
    fix P Q
    assume "(P, Q)  symcl (indRelRST SRel TRel)"
    thus "P ℛ⟦⋅⟧<SRel,TRel> Q"
    proof (auto simp add: symcl_def indRelRST.simps)
      fix S
      show "SourceTerm S ℛ⟦⋅⟧<SRel,TRel> TargetTerm (S)"
        by (rule indRelST.encR)
    next
      fix S1 S2
      assume "(S1, S2)  SRel"
      thus "SourceTerm S1 ℛ⟦⋅⟧<SRel,TRel> SourceTerm S2"
        by (rule indRelST.source)
    next
      fix T1 T2
      assume "(T1, T2)  TRel"
      thus "TargetTerm T1 ℛ⟦⋅⟧<SRel,TRel> TargetTerm T2"
        by (rule indRelST.target)
    next
      fix S
      show "TargetTerm (S) ℛ⟦⋅⟧<SRel,TRel> SourceTerm S"
        by (rule indRelST.encL)
    next
      fix S1 S2
      assume "(S1, S2)  SRel"
      with symmS show "SourceTerm S2 ℛ⟦⋅⟧<SRel,TRel> SourceTerm S1"
          unfolding sym_def
        by (simp add: indRelST.source)
    next
      fix T1 T2
      assume "(T1, T2)  TRel"
      with symmT show "(TargetTerm T2, TargetTerm T1)  indRelST SRel TRel"
          unfolding sym_def
        by (simp add: indRelST.target)
    qed
  qed
  with reflT show "indRelSTEQ SRel TRel = (symcl ((indRelRST SRel TRel)=))+"
      using refl_symm_trans_closure_is_symm_refl_trans_closure[where Rel="indRelRST SRel TRel"]
            refl_trans_closure_of_indRelST
    by simp
qed

lemma (in encoding) symm_closure_of_indRelLST:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes reflT: "refl TRel"
      and symmS: "sym SRel"
      and symmT: "sym TRel"
  shows "indRelST SRel TRel = symcl (indRelLST SRel TRel)"
    and "indRelSTEQ SRel TRel = (symcl ((indRelLST SRel TRel)=))+"
proof -
  show "indRelST SRel TRel = symcl (indRelLST SRel TRel)"
  proof auto
    fix P Q
    assume "P ℛ⟦⋅⟧<SRel,TRel> Q"
    thus "(P, Q)  symcl (indRelLST SRel TRel)"
      by (induct, simp_all add: symcl_def indRelLST.intros)
  next
    fix P Q
    assume "(P, Q)  symcl (indRelLST SRel TRel)"
    thus "P ℛ⟦⋅⟧<SRel,TRel> Q"
    proof (auto simp add: symcl_def indRelLST.simps)
      fix S
      show "SourceTerm S ℛ⟦⋅⟧<SRel,TRel> TargetTerm (S)"
        by (rule indRelST.encR)
    next
      fix S1 S2
      assume "(S1, S2)  SRel"
      thus "SourceTerm S1 ℛ⟦⋅⟧<SRel,TRel> SourceTerm S2"
        by (rule indRelST.source)
    next
      fix T1 T2
      assume "(T1, T2)  TRel"
      thus "TargetTerm T1 ℛ⟦⋅⟧<SRel,TRel> TargetTerm T2"
        by (rule indRelST.target)
    next
      fix S
      show "TargetTerm (S) ℛ⟦⋅⟧<SRel,TRel> SourceTerm S"
        by (rule indRelST.encL)
    next
      fix S1 S2
      assume "(S1, S2)  SRel"
      with symmS show "SourceTerm S2 ℛ⟦⋅⟧<SRel,TRel> SourceTerm S1"
          unfolding sym_def
        by (simp add: indRelST.source)
    next
      fix T1 T2
      assume "(T1, T2)  TRel"
      with symmT show "TargetTerm T2 ℛ⟦⋅⟧<SRel,TRel> TargetTerm T1"
          unfolding sym_def
        by (simp add: indRelST.target)
    qed
  qed
  with reflT show "indRelSTEQ SRel TRel = (symcl ((indRelLST SRel TRel)=))+"
      using refl_symm_trans_closure_is_symm_refl_trans_closure[where Rel="indRelLST SRel TRel"]
            refl_trans_closure_of_indRelST
    by simp
qed

lemma (in encoding) symm_trans_closure_of_indRelRSTPO:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes symmS: "sym SRel"
      and symmT: "sym TRel"
  shows "indRelSTEQ SRel TRel = (symcl (indRelRSTPO SRel TRel))+"
proof auto
  fix P Q
  assume "P ∼⟦⋅⟧<SRel,TRel> Q"
  thus "(P, Q)  (symcl (indRelRSTPO SRel TRel))+"
  proof induct
    case (encR S)
    show "(SourceTerm S, TargetTerm (S))  (symcl (indRelRSTPO SRel TRel))+"
        using indRelRSTPO.encR[of S SRel TRel]
        unfolding symcl_def
      by auto
  next
    case (encL S)
    show "(TargetTerm (S), SourceTerm S)  (symcl (indRelRSTPO SRel TRel))+"
        using indRelRSTPO.encR[of S SRel TRel]
        unfolding symcl_def
      by auto
  next
    case (source S1 S2)
    assume "(S1, S2)  SRel"
    thus "(SourceTerm S1, SourceTerm S2)  (symcl (indRelRSTPO SRel TRel))+"
        using indRelRSTPO.source[of S1 S2 SRel TRel]
        unfolding symcl_def
      by auto
  next
    case (target T1 T2)
    assume "(T1, T2)  TRel"
    thus "(TargetTerm T1, TargetTerm T2)  (symcl (indRelRSTPO SRel TRel))+"
        using indRelRSTPO.target[of T1 T2 TRel SRel]
        unfolding symcl_def
      by auto
  next
    case (trans P Q R)
    assume "(P, Q)  (symcl (indRelRSTPO SRel TRel))+"
       and "(Q, R)  (symcl (indRelRSTPO SRel TRel))+"
    thus "(P, R)  (symcl (indRelRSTPO SRel TRel))+"
      by simp
  qed
next
  fix P Q
  assume "(P, Q)  (symcl (indRelRSTPO SRel TRel))+"
  thus "P ∼⟦⋅⟧<SRel,TRel> Q"
  proof induct
    fix Q
    assume "(P, Q)  symcl (indRelRSTPO SRel TRel)"
    thus "P ∼⟦⋅⟧<SRel,TRel> Q"
    proof (cases "P ≲⟦⋅⟧R<SRel,TRel> Q", simp_all add: symcl_def)
      assume "P ≲⟦⋅⟧R<SRel,TRel> Q"
      thus "P ∼⟦⋅⟧<SRel,TRel> Q"
      proof induct
        case (encR S)
        show "SourceTerm S ∼⟦⋅⟧<SRel,TRel> TargetTerm (S)"
          by (rule indRelSTEQ.encR)
      next
        case (source S1 S2)
        assume "(S1, S2)  SRel"
        thus "SourceTerm S1 ∼⟦⋅⟧<SRel,TRel> SourceTerm S2"
          by (rule indRelSTEQ.source)
      next
        case (target T1 T2)
        assume "(T1, T2)  TRel"
        thus "TargetTerm T1 ∼⟦⋅⟧<SRel,TRel> TargetTerm T2"
          by (rule indRelSTEQ.target)
      next
        case (trans P Q R)
        assume "P ∼⟦⋅⟧<SRel,TRel> Q" and "Q ∼⟦⋅⟧<SRel,TRel> R"
        thus "P ∼⟦⋅⟧<SRel,TRel> R"
          by (rule indRelSTEQ.trans)
      qed
    next
      assume "Q ≲⟦⋅⟧R<SRel,TRel> P"
      thus "P ∼⟦⋅⟧<SRel,TRel> Q"
      proof induct
        case (encR S)
        show "TargetTerm (S) ∼⟦⋅⟧<SRel,TRel> SourceTerm S"
          by (rule indRelSTEQ.encL)
      next
        case (source S1 S2)
        assume "(S1, S2)  SRel"
        with symmS show "SourceTerm S2 ∼⟦⋅⟧<SRel,TRel> SourceTerm S1"
            unfolding sym_def
          by (simp add: indRelSTEQ.source)
      next
        case (target T1 T2)
        assume "(T1, T2)  TRel"
        with symmT show "TargetTerm T2 ∼⟦⋅⟧<SRel,TRel> TargetTerm T1"
            unfolding sym_def
          by (simp add: indRelSTEQ.target)
      next
        case (trans P Q R)
        assume "R ∼⟦⋅⟧<SRel,TRel> Q" and "Q ∼⟦⋅⟧<SRel,TRel> P"
        thus "R ∼⟦⋅⟧<SRel,TRel> P"
          by (rule indRelSTEQ.trans)
      qed
    qed
  next
    case (step Q R)
    assume "P ∼⟦⋅⟧<SRel,TRel> Q"
    moreover assume "(Q, R)  symcl (indRelRSTPO SRel TRel)"
    hence "Q ∼⟦⋅⟧<SRel,TRel> R"
    proof (auto simp add: symcl_def)
      assume "Q ≲⟦⋅⟧R<SRel,TRel> R"
      thus "Q ∼⟦⋅⟧<SRel,TRel> R"
      proof (induct, simp add: indRelSTEQ.encR, simp add: indRelSTEQ.source,
             simp add: indRelSTEQ.target)
        case (trans P Q R)
        assume "P ∼⟦⋅⟧<SRel,TRel> Q" and "Q ∼⟦⋅⟧<SRel,TRel> R"
        thus "P ∼⟦⋅⟧<SRel,TRel> R"
          by (rule indRelSTEQ.trans)
      qed
    next
      assume "R ≲⟦⋅⟧R<SRel,TRel> Q"
      hence "R ∼⟦⋅⟧<SRel,TRel> Q"
      proof (induct, simp add: indRelSTEQ.encR, simp add: indRelSTEQ.source,
             simp add: indRelSTEQ.target)
        case (trans P Q R)
        assume "P ∼⟦⋅⟧<SRel,TRel> Q" and "Q ∼⟦⋅⟧<SRel,TRel> R"
        thus "P ∼⟦⋅⟧<SRel,TRel> R"
          by (rule indRelSTEQ.trans)
      qed
      with symmS symmT show "Q ∼⟦⋅⟧<SRel,TRel> R"
          using indRelSTEQ_symm[of SRel TRel]
          unfolding sym_def
        by blast
    qed
    ultimately show "P ∼⟦⋅⟧<SRel,TRel> R"
      by (rule indRelSTEQ.trans)
  qed
qed

lemma (in encoding) symm_trans_closure_of_indRelLSTPO:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes symmS: "sym SRel"
      and symmT: "sym TRel"
  shows "indRelSTEQ SRel TRel = (symcl (indRelLSTPO SRel TRel))+"
proof auto
  fix P Q
  assume "P ∼⟦⋅⟧<SRel,TRel> Q"
  thus "(P, Q)  (symcl (indRelLSTPO SRel TRel))+"
  proof induct
    case (encR S)
    show "(SourceTerm S, TargetTerm (S))  (symcl (indRelLSTPO SRel TRel))+"
        using indRelLSTPO.encL[of S SRel TRel]
        unfolding symcl_def
      by blast
  next
    case (encL S)
    show "(TargetTerm (S), SourceTerm S)  (symcl (indRelLSTPO SRel TRel))+"
        using indRelLSTPO.encL[of S SRel TRel]
        unfolding symcl_def
      by blast
  next
    case (source S1 S2)
    assume "(S1, S2)  SRel"
    thus "(SourceTerm S1, SourceTerm S2)  (symcl (indRelLSTPO SRel TRel))+"
        using indRelLSTPO.source[of S1 S2 SRel TRel]
        unfolding symcl_def
      by blast
  next
    case (target T1 T2)
    assume "(T1, T2)  TRel"
    thus "(TargetTerm T1, TargetTerm T2)  (symcl (indRelLSTPO SRel TRel))+"
        using indRelLSTPO.target[of T1 T2 TRel SRel]
        unfolding symcl_def
      by blast
  next
    case (trans P Q R)
    assume "(P, Q)  (symcl (indRelLSTPO SRel TRel))+"
       and "(Q, R)  (symcl (indRelLSTPO SRel TRel))+"
    thus "(P, R)  (symcl (indRelLSTPO SRel TRel))+"
      by simp
  qed
next
  fix P Q
  assume "(P, Q)  (symcl (indRelLSTPO SRel TRel))+"
  thus "P ∼⟦⋅⟧<SRel,TRel> Q"
  proof induct
    fix Q
    assume "(P, Q)  symcl (indRelLSTPO SRel TRel)"
    thus "P ∼⟦⋅⟧<SRel,TRel> Q"
      unfolding symcl_def
    proof auto
      assume "P ≲⟦⋅⟧L<SRel,TRel> Q"
      thus "P ∼⟦⋅⟧<SRel,TRel> Q"
      proof (induct, simp add: indRelSTEQ.encL, simp add: indRelSTEQ.source,
             simp add: indRelSTEQ.target)
        case (trans P Q R)
        assume "P ∼⟦⋅⟧<SRel,TRel> Q" and "Q ∼⟦⋅⟧<SRel,TRel> R"
        thus "P ∼⟦⋅⟧<SRel,TRel> R"
          by (rule indRelSTEQ.trans)
      qed
    next
      assume "Q ≲⟦⋅⟧L<SRel,TRel> P"
      hence "Q ∼⟦⋅⟧<SRel,TRel> P"
      proof (induct, simp add: indRelSTEQ.encL, simp add: indRelSTEQ.source,
             simp add: indRelSTEQ.target)
        case (trans P Q R)
        assume "P ∼⟦⋅⟧<SRel,TRel> Q" and "Q ∼⟦⋅⟧<SRel,TRel> R"
        thus "P ∼⟦⋅⟧<SRel,TRel> R"
          by (rule indRelSTEQ.trans)
      qed
      with symmS symmT show "P ∼⟦⋅⟧<SRel,TRel> Q"
          using indRelSTEQ_symm[of SRel TRel]
          unfolding sym_def
        by blast
    qed
  next
    case (step Q R)
    assume "P ∼⟦⋅⟧<SRel,TRel> Q"
    moreover assume "(Q, R)  symcl (indRelLSTPO SRel TRel)"
    hence "Q ∼⟦⋅⟧<SRel,TRel> R"
      unfolding symcl_def
    proof auto
      assume "Q ≲⟦⋅⟧L<SRel,TRel> R"
      thus "Q ∼⟦⋅⟧<SRel,TRel> R"
      proof (induct, simp add: indRelSTEQ.encL, simp add: indRelSTEQ.source,
             simp add: indRelSTEQ.target)
        case (trans P Q R)
        assume "P ∼⟦⋅⟧<SRel,TRel> Q" and "Q ∼⟦⋅⟧<SRel,TRel> R"
        thus "P ∼⟦⋅⟧<SRel,TRel> R"
          by (rule indRelSTEQ.trans)
      qed
    next
      assume "R ≲⟦⋅⟧L<SRel,TRel> Q"
      hence "R ∼⟦⋅⟧<SRel,TRel> Q"
      proof (induct, simp add: indRelSTEQ.encL, simp add: indRelSTEQ.source,
             simp add: indRelSTEQ.target)
        case (trans P Q R)
        assume "P ∼⟦⋅⟧<SRel,TRel> Q" and "Q ∼⟦⋅⟧<SRel,TRel> R"
        thus "P ∼⟦⋅⟧<SRel,TRel> R"
          by (rule indRelSTEQ.trans)
      qed
      with symmS symmT show "Q ∼⟦⋅⟧<SRel,TRel> R"
          using indRelSTEQ_symm[of SRel TRel]
          unfolding sym_def
        by blast
    qed
    ultimately show "P ∼⟦⋅⟧<SRel,TRel> R"
      by (rule indRelSTEQ.trans)
  qed
qed

text ‹If the relations indRelRST, indRelLST, or indRelST contain a pair of target terms, then
        this pair is also related by the considered target term relation. Similarly a pair of
        source terms is related by the considered source term relation.›

lemma (in encoding) indRelRST_to_SRel:
  fixes SRel  :: "('procS × 'procS) set"
    and TRel  :: "('procT × 'procT) set"
    and SP SQ :: "'procS"
  assumes rel: "SourceTerm SP ℛ⟦⋅⟧R<SRel,TRel> SourceTerm SQ"
  shows "(SP, SQ)  SRel"
      using rel
    by (simp add: indRelRST.simps)

lemma (in encoding) indRelRST_to_TRel:
  fixes SRel  :: "('procS × 'procS) set"
    and TRel  :: "('procT × 'procT) set"
    and TP TQ :: "'procT"
  assumes rel: "TargetTerm TP ℛ⟦⋅⟧R<SRel,TRel> TargetTerm TQ"
  shows "(TP, TQ)  TRel"
      using rel
    by (simp add: indRelRST.simps)

lemma (in encoding) indRelLST_to_SRel:
  fixes SRel  :: "('procS × 'procS) set"
    and TRel  :: "('procT × 'procT) set"
    and SP SQ :: "'procS"
  assumes rel: "SourceTerm SP ℛ⟦⋅⟧L<SRel,TRel> SourceTerm SQ"
  shows "(SP, SQ)  SRel"
      using rel
    by (simp add: indRelLST.simps)

lemma (in encoding) indRelLST_to_TRel:
  fixes SRel  :: "('procS × 'procS) set"
    and TRel  :: "('procT × 'procT) set"
    and TP TQ :: "'procT"
  assumes rel: "TargetTerm TP ℛ⟦⋅⟧L<SRel,TRel> TargetTerm TQ"
  shows "(TP, TQ)  TRel"
      using rel
    by (simp add: indRelLST.simps)

lemma (in encoding) indRelST_to_SRel:
  fixes SRel  :: "('procS × 'procS) set"
    and TRel  :: "('procT × 'procT) set"
    and SP SQ :: "'procS"
  assumes rel: "SourceTerm SP ℛ⟦⋅⟧<SRel,TRel> SourceTerm SQ"
  shows "(SP, SQ)  SRel"
      using rel
    by (simp add: indRelST.simps)

lemma (in encoding) indRelST_to_TRel:
  fixes SRel  :: "('procS × 'procS) set"
    and TRel  :: "('procT × 'procT) set"
    and TP TQ :: "'procT"
  assumes rel: "TargetTerm TP ℛ⟦⋅⟧<SRel,TRel> TargetTerm TQ"
  shows "(TP, TQ)  TRel"
      using rel
    by (simp add: indRelST.simps)

text ‹If the relations indRelRSTPO or indRelLSTPO contain a pair of target terms, then this pair
        is also related by the transitive closure of the considered target term relation. Similarly
        a pair of source terms is related by the transitive closure of the source term relation. A
        pair of a source and a target term results from the combination of pairs in the source
        relation, the target relation, and the encoding function. Note that, because of the
        symmetry, no similar condition holds for indRelSTEQ.›

lemma (in encoding) indRelRSTPO_to_SRel_and_TRel:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
    and P Q  :: "('procS, 'procT) Proc"
  assumes "P ≲⟦⋅⟧R<SRel,TRel> Q"
  shows "SP SQ. SP ∈S P  SQ ∈S Q  (SP, SQ)  SRel+"
    and "SP TQ. SP ∈S P  TQ ∈T Q  (S. (SP, S)  SRel*  (S, TQ)  TRel*)"
    and "TP SQ. TP ∈T P  SQ ∈S Q  False"
    and "TP TQ. TP ∈T P  TQ ∈T Q  (TP, TQ)  TRel+"
      using assms
proof induct
  case (encR S)
  show "SP SQ. SP ∈S SourceTerm S  SQ ∈S TargetTerm (S)  (SP, SQ)  SRel+"
   and "TP SQ. TP ∈T SourceTerm S  SQ ∈S TargetTerm (S)  False"
   and "TP TQ. TP ∈T SourceTerm S  TQ ∈T TargetTerm (S)  (TP, TQ)  TRel+"
    by simp+
  have "(S, S)  SRel*"
    by simp
  moreover have "(S, S)  TRel*"
    by simp
  ultimately show "SP TQ. SP ∈S SourceTerm S  TQ ∈T TargetTerm (S) 
                   (S. (SP, S)  SRel*  (S, TQ)  TRel*)"
    by blast
next
  case (source S1 S2)
  assume "(S1, S2)  SRel"
  thus "SP SQ. SP ∈S SourceTerm S1  SQ ∈S SourceTerm S2  (SP, SQ)  SRel+"
    by simp
  show "SP TQ. SP ∈S SourceTerm S1  TQ ∈T SourceTerm S2 
        (S. (SP, S)  SRel*  (S, TQ)  TRel*)"
   and "TP SQ. TP ∈T SourceTerm S1  SQ ∈S SourceTerm S2  False"
   and "TP TQ. TP ∈T SourceTerm S1  TQ ∈T SourceTerm S2  (TP, TQ)  TRel+"
    by simp+
next
  case (target T1 T2)
  show "SP SQ. SP ∈S TargetTerm T1  SQ ∈S TargetTerm T2  (SP, SQ)  SRel+"
   and "SP TQ. SP ∈S TargetTerm T1  TQ ∈T TargetTerm T2
         (S. (SP, S)  SRel*  (S, TQ)  TRel*)"
   and "TP SQ. TP ∈T TargetTerm T1  SQ ∈S TargetTerm T2  False"
    by simp+
  assume "(T1, T2)  TRel"
  thus "TP TQ. TP ∈T TargetTerm T1  TQ ∈T TargetTerm T2  (TP, TQ)  TRel+"
    by simp
next
  case (trans P Q R)
  assume A1: "SP SQ. SP ∈S P  SQ ∈S Q  (SP, SQ)  SRel+"
     and A2: "SP TQ. SP ∈S P  TQ ∈T Q  (S. (SP, S)  SRel*  (S, TQ)  TRel*)"
     and A3: "TP SQ. TP ∈T P  SQ ∈S Q  False"
     and A4: "TP TQ. TP ∈T P  TQ ∈T Q  (TP, TQ)  TRel+"
     and A5: "SQ SR. SQ ∈S Q  SR ∈S R  (SQ, SR)  SRel+"
     and A6: "SQ TR. SQ ∈S Q  TR ∈T R  (S. (SQ, S)  SRel*  (S, TR)  TRel*)"
     and A7: "TQ SR. TQ ∈T Q  SR ∈S R  False"
     and A8: "TQ TR. TQ ∈T Q  TR ∈T R  (TQ, TR)  TRel+"
  show "SP SR. SP ∈S P  SR ∈S R  (SP, SR)  SRel+"
  proof clarify
    fix SP SR
    assume A9: "SP ∈S P" and A10: "SR ∈S R"
    show "(SP, SR)  SRel+"
    proof (cases Q)
      case (SourceTerm SQ)
      assume A11: "SQ ∈S Q"
      with A1 A9 have "(SP, SQ)  SRel+"
        by simp
      moreover from A5 A10 A11 have "(SQ, SR)  SRel+"
        by simp
      ultimately show "(SP, SR)  SRel+"
        by simp
    next
      case (TargetTerm TQ)
      assume "TQ ∈T Q"
      with A7 A10 show "(SP, SR)  SRel+"
        by blast
    qed
  qed
  show "SP TR. SP ∈S P  TR ∈T R
         (S. (SP, S)  SRel*  (S, TR)  TRel*)"
  proof clarify
    fix SP TR
    assume A9: "SP ∈S P" and A10: "TR ∈T R"
    show "S. (SP, S)  SRel*  (S, TR)  TRel*"
    proof (cases Q)
      case (SourceTerm SQ)
      assume A11: "SQ ∈S Q"
      with A6 A10 obtain S where A12: "(SQ, S)  SRel*"
                             and A13: "(S, TR)  TRel*"
        by blast
      from A1 A9 A11 have "(SP, SQ)  SRel*"
        by simp
      from this A12 have "(SP, S)  SRel*"
        by simp
      with A13 show "S. (SP, S)  SRel*  (S, TR)  TRel*"
        by blast
    next
      case (TargetTerm TQ)
      assume A11: "TQ ∈T Q"
      with A2 A9 obtain S where A12: "(SP, S)  SRel*"
                            and A13: "(S, TQ)  TRel*"
        by blast
      from A8 A10 A11 have "(TQ, TR)  TRel*"
        by simp
      with A13 have "(S, TR)  TRel*"
        by simp
      with A12 show "S. (SP, S)  SRel*  (S, TR)  TRel*"
        by blast
    qed
  qed
  show "TP SR. TP ∈T P  SR ∈S R  False"
  proof clarify
    fix TP SR
    assume A9: "TP ∈T P" and A10: "SR ∈S R"
    show "False"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A3 A9 show "False"
        by blast
    next
      case (TargetTerm TQ)
      assume "TQ ∈T Q"
      with A7 A10 show "False"
        by blast
    qed
  qed
  show "TP TR. TP ∈T P  TR ∈T R  (TP, TR)  TRel+"
  proof clarify
    fix TP TR
    assume A9: "TP ∈T P" and A10: "TR ∈T R"
    show "(TP, TR)  TRel+"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A3 A9 show "(TP, TR)  TRel+"
        by blast
    next
      case (TargetTerm TQ)
      assume A11: "TQ ∈T Q"
      with A4 A9 have "(TP, TQ)  TRel+"
        by simp
      moreover from A8 A10 A11 have "(TQ, TR)  TRel+"
        by simp
      ultimately show "(TP, TR)  TRel+"
        by simp
    qed
  qed
qed

lemma (in encoding) indRelLSTPO_to_SRel_and_TRel:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
    and P Q  :: "('procS, 'procT) Proc"
  assumes "P ≲⟦⋅⟧L<SRel,TRel> Q"
  shows "SP SQ. SP ∈S P  SQ ∈S Q  (SP, SQ)  SRel+"
    and "SP TQ. SP ∈S P  TQ ∈T Q  False"
    and "TP SQ. TP ∈T P  SQ ∈S Q  (S. (TP, S)  TRel*  (S, SQ)  SRel*)"
    and "TP TQ. TP ∈T P  TQ ∈T Q  (TP, TQ)  TRel+"
      using assms
proof induct
  case (encL S)
  show "SP SQ. SP ∈S TargetTerm (S)  SQ ∈S SourceTerm S  (SP, SQ)  SRel+"
   and "SP TQ. SP ∈S TargetTerm (S)  TQ ∈T SourceTerm S  False"
   and "TP TQ. TP ∈T TargetTerm (S)  TQ ∈T SourceTerm S  (TP, TQ)  TRel+"
    by simp+
  have "(S, S)  TRel*"
    by simp
  moreover have "(S, S)  SRel*"
    by simp
  ultimately show "TP SQ. TP ∈T TargetTerm (S)  SQ ∈S SourceTerm S 
                   (S. (TP, S)  TRel*  (S, SQ)  SRel*)"
    by blast
next
  case (source S1 S2)
  assume "(S1, S2)  SRel"
  thus "SP SQ. SP ∈S SourceTerm S1  SQ ∈S SourceTerm S2  (SP, SQ)  SRel+"
    by simp
  show "SP TQ. SP ∈S SourceTerm S1  TQ ∈T SourceTerm S2  False"
   and "TP SQ. TP ∈T SourceTerm S1  SQ ∈S SourceTerm S2
         (S. (TP, S)  TRel*  (S, SQ)  SRel*)"
   and "TP TQ. TP ∈T SourceTerm S1  TQ ∈T SourceTerm S2  (TP, TQ)  TRel+"
    by simp+
next
  case (target T1 T2)
  show "SP SQ. SP ∈S TargetTerm T1  SQ ∈S TargetTerm T2  (SP, SQ)  SRel+"
   and "SP TQ. SP ∈S TargetTerm T1  TQ ∈T TargetTerm T2  False"
   and "TP SQ. TP ∈T TargetTerm T1  SQ ∈S TargetTerm T2
         (S. (TP, S)  TRel*  (S, SQ)  SRel*)"
    by simp+
  assume "(T1, T2)  TRel"
  thus "TP TQ. TP ∈T TargetTerm T1  TQ ∈T TargetTerm T2  (TP, TQ)  TRel+"
    by simp
next
  case (trans P Q R)
  assume A1: "SP SQ. SP ∈S P  SQ ∈S Q  (SP, SQ)  SRel+"
     and A2: "SP TQ. SP ∈S P  TQ ∈T Q  False"
     and A3: "TP SQ. TP ∈T P  SQ ∈S Q
               (S. (TP, S)  TRel*  (S, SQ)  SRel*)"
     and A4: "TP TQ. TP ∈T P  TQ ∈T Q  (TP, TQ)  TRel+"
     and A5: "SQ SR. SQ ∈S Q  SR ∈S R  (SQ, SR)  SRel+"
     and A6: "SQ TR. SQ ∈S Q  TR ∈T R  False"
     and A7: "TQ SR. TQ ∈T Q  SR ∈S R  (S. (TQ, S)  TRel*  (S, SR)  SRel*)"
     and A8: "TQ TR. TQ ∈T Q  TR ∈T R  (TQ, TR)  TRel+"
  show "SP SR. SP ∈S P  SR ∈S R  (SP, SR)  SRel+"
  proof clarify
    fix SP SR
    assume A9: "SP ∈S P" and A10: "SR ∈S R"
    show "(SP, SR)  SRel+"
    proof (cases Q)
      case (SourceTerm SQ)
      assume A11: "SQ ∈S Q"
      with A1 A9 have "(SP, SQ)  SRel+"
        by simp
      moreover from A5 A10 A11 have "(SQ, SR)  SRel+"
        by simp
      ultimately show "(SP, SR)  SRel+"
        by simp
    next
      case (TargetTerm TQ)
      assume "TQ ∈T Q"
      with A2 A9 show "(SP, SR)  SRel+"
        by blast
    qed
  qed
  show "SP TR. SP ∈S P  TR ∈T R  False"
  proof clarify
    fix SP TR
    assume A9: "SP ∈S P" and A10: "TR ∈T R"
    show "False"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A6 A10 show "False"
        by blast
    next
      case (TargetTerm TQ)
      assume "TQ ∈T Q"
      with A2 A9 show "False"
        by blast
    qed
  qed
  show "TP SR. TP ∈T P  SR ∈S R  (S. (TP, S)  TRel*  (S, SR)  SRel*)"
  proof clarify
    fix TP SR
    assume A9: "TP ∈T P" and A10: "SR ∈S R"
    show "S. (TP, S)  TRel*  (S, SR)  SRel*"
    proof (cases Q)
      case (SourceTerm SQ)
      assume A11: "SQ ∈S Q"
      with A3 A9 obtain S where A12: "(TP, S)  TRel*" and A13: "(S, SQ)  SRel*"
        by blast
      from A5 A10 A11 have "(SQ, SR)  SRel*"
        by simp
      with A13 have "(S, SR)  SRel*"
        by simp
      with A12 show "S. (TP, S)  TRel*  (S, SR)  SRel*"
        by blast
    next
      case (TargetTerm TQ)
      assume A11: "TQ ∈T Q"
      with A7 A10 obtain S where A12: "(TQ, S)  TRel*" and A13: "(S, SR)  SRel*"
        by blast
      from A4 A9 A11 have "(TP, TQ)  TRel*"
        by simp
      from this A12 have "(TP, S)  TRel*"
        by simp
      with A13 show "S. (TP, S)  TRel*  (S, SR)  SRel*"
        by blast
    qed
  qed
  show "TP TR. TP ∈T P  TR ∈T R  (TP, TR)  TRel+"
  proof clarify
    fix TP TR
    assume A9: "TP ∈T P" and A10: "TR ∈T R"
    show "(TP, TR)  TRel+"
    proof (cases Q)
      case (SourceTerm SQ)
      assume "SQ ∈S Q"
      with A6 A10 show "(TP, TR)  TRel+"
        by blast
    next
      case (TargetTerm TQ)
      assume A11: "TQ ∈T Q"
      with A4 A9 have "(TP, TQ)  TRel+"
        by simp
      moreover from A8 A10 A11 have "(TQ, TR)  TRel+"
        by simp
      ultimately show "(TP, TR)  TRel+"
        by simp
    qed
  qed
qed

text ‹If indRelRSTPO, indRelLSTPO, or indRelSTPO preserves barbs then so do the corresponding
        source term and target term relations.›

lemma (in encoding_wrt_barbs) rel_with_source_impl_SRel_preserves_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes preservation: "rel_preserves_barbs Rel (STCalWB SWB TWB)"
      and sourceInRel:  "S1 S2. (S1, S2)  SRel  (SourceTerm S1, SourceTerm S2)  Rel"
  shows "rel_preserves_barbs SRel SWB"
proof clarify
  fix SP SQ a
  assume "(SP, SQ)  SRel"
  with sourceInRel have "(SourceTerm SP, SourceTerm SQ)  Rel"
    by blast
  moreover assume "SP↓<SWB>a"
  hence "SourceTerm SP↓.a"
    by simp
  ultimately have "SourceTerm SQ↓.a"
      using preservation preservation_of_barbs_in_barbed_encoding[where Rel="Rel"]
    by blast
  thus "SQ↓<SWB>a"
    by simp
qed

lemma (in encoding_wrt_barbs) indRelRSTPO_impl_SRel_and_TRel_preserve_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes preservation: "rel_preserves_barbs (indRelRSTPO SRel TRel) (STCalWB SWB TWB)"
  shows "rel_preserves_barbs SRel SWB"
    and "rel_preserves_barbs TRel TWB"
proof -
  show "rel_preserves_barbs SRel SWB"
      using preservation rel_with_source_impl_SRel_preserves_barbs[where
                          Rel="indRelRSTPO SRel TRel" and SRel="SRel"]
    by (simp add: indRelRSTPO.source)
next
  show "rel_preserves_barbs TRel TWB"
      using preservation rel_with_target_impl_TRel_preserves_barbs[where
                          Rel="indRelRSTPO SRel TRel" and TRel="TRel"]
    by (simp add: indRelRSTPO.target)
qed

lemma (in encoding_wrt_barbs) indRelLSTPO_impl_SRel_and_TRel_preserve_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes preservation: "rel_preserves_barbs (indRelLSTPO SRel TRel) (STCalWB SWB TWB)"
  shows "rel_preserves_barbs SRel SWB"
    and "rel_preserves_barbs TRel TWB"
proof -
  show "rel_preserves_barbs SRel SWB"
      using preservation rel_with_source_impl_SRel_preserves_barbs[where
                          Rel="indRelLSTPO SRel TRel" and SRel="SRel"]
    by (simp add: indRelLSTPO.source)
next
  show "rel_preserves_barbs TRel TWB"
      using preservation rel_with_target_impl_TRel_preserves_barbs[where
                          Rel="indRelLSTPO SRel TRel" and TRel="TRel"]
    by (simp add: indRelLSTPO.target)
qed

lemma (in encoding_wrt_barbs) indRelSTEQ_impl_SRel_and_TRel_preserve_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes preservation: "rel_preserves_barbs (indRelSTEQ SRel TRel) (STCalWB SWB TWB)"
  shows "rel_preserves_barbs SRel SWB"
    and "rel_preserves_barbs TRel TWB"
proof -
  show "rel_preserves_barbs SRel SWB"
      using preservation rel_with_source_impl_SRel_preserves_barbs[where
                          Rel="indRelSTEQ SRel TRel" and SRel="SRel"]
    by (simp add: indRelSTEQ.source)
next
  show "rel_preserves_barbs TRel TWB"
      using preservation rel_with_target_impl_TRel_preserves_barbs[where
                          Rel="indRelSTEQ SRel TRel" and TRel="TRel"]
    by (simp add: indRelSTEQ.target)
qed

lemma (in encoding_wrt_barbs) rel_with_source_impl_SRel_weakly_preserves_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes preservation: "rel_weakly_preserves_barbs Rel (STCalWB SWB TWB)"
      and sourceInRel:  "S1 S2. (S1, S2)  SRel  (SourceTerm S1, SourceTerm S2)  Rel"
  shows "rel_weakly_preserves_barbs SRel SWB"
proof clarify
  fix SP SQ a SP'
  assume "(SP, SQ)  SRel"
  with sourceInRel have "(SourceTerm SP, SourceTerm SQ)  Rel"
    by blast
  moreover assume "SP (Calculus SWB)* SP'" and "SP'↓<SWB>a"
  hence "SourceTerm SP⇓.a"
    by blast
  ultimately have "SourceTerm SQ⇓.a"
      using preservation weak_preservation_of_barbs_in_barbed_encoding[where Rel="Rel"]
    by blast
  thus "SQ⇓<SWB>a"
    by simp
qed

lemma (in encoding_wrt_barbs) indRelRSTPO_impl_SRel_and_TRel_weakly_preserve_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes preservation: "rel_weakly_preserves_barbs (indRelRSTPO SRel TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_preserves_barbs SRel SWB"
    and "rel_weakly_preserves_barbs TRel TWB"
proof -
  show "rel_weakly_preserves_barbs SRel SWB"
      using preservation rel_with_source_impl_SRel_weakly_preserves_barbs[where
                          Rel="indRelRSTPO SRel TRel" and SRel="SRel"]
    by (simp add: indRelRSTPO.source)
next
  show "rel_weakly_preserves_barbs TRel TWB"
      using preservation rel_with_target_impl_TRel_weakly_preserves_barbs[where
                          Rel="indRelRSTPO SRel TRel" and TRel="TRel"]
    by (simp add: indRelRSTPO.target)
qed

lemma (in encoding_wrt_barbs) indRelLSTPO_impl_SRel_and_TRel_weakly_preserve_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes preservation: "rel_weakly_preserves_barbs (indRelLSTPO SRel TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_preserves_barbs SRel SWB"
    and "rel_weakly_preserves_barbs TRel TWB"
proof -
  show "rel_weakly_preserves_barbs SRel SWB"
      using preservation rel_with_source_impl_SRel_weakly_preserves_barbs[where
                          Rel="indRelLSTPO SRel TRel" and SRel="SRel"]
    by (simp add: indRelLSTPO.source)
next
  show "rel_weakly_preserves_barbs TRel TWB"
      using preservation rel_with_target_impl_TRel_weakly_preserves_barbs[where
                          Rel="indRelLSTPO SRel TRel" and TRel="TRel"]
    by (simp add: indRelLSTPO.target)
qed

lemma (in encoding_wrt_barbs) indRelSTEQ_impl_SRel_and_TRel_weakly_preserve_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes preservation: "rel_weakly_preserves_barbs (indRelSTEQ SRel TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_preserves_barbs SRel SWB"
    and "rel_weakly_preserves_barbs TRel TWB"
proof -
  show "rel_weakly_preserves_barbs SRel SWB"
      using preservation rel_with_source_impl_SRel_weakly_preserves_barbs[where
                          Rel="indRelSTEQ SRel TRel" and SRel="SRel"]
    by (simp add: indRelSTEQ.source)
next
  show "rel_weakly_preserves_barbs TRel TWB"
      using preservation rel_with_target_impl_TRel_weakly_preserves_barbs[where
                          Rel="indRelSTEQ SRel TRel" and TRel="TRel"]
    by (simp add: indRelSTEQ.target)
qed

text ‹If indRelRSTPO, indRelLSTPO, or indRelSTPO reflects barbs then so do the corresponding
        source term and target term relations.›

lemma (in encoding_wrt_barbs) rel_with_source_impl_SRel_reflects_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes reflection: "rel_reflects_barbs Rel (STCalWB SWB TWB)"
      and sourceInRel:  "S1 S2. (S1, S2)  SRel  (SourceTerm S1, SourceTerm S2)  Rel"
  shows "rel_reflects_barbs SRel SWB"
proof clarify
  fix SP SQ a
  assume "(SP, SQ)  SRel"
  with sourceInRel have "(SourceTerm SP, SourceTerm SQ)  Rel"
    by blast
  moreover assume "SQ↓<SWB>a"
  hence "SourceTerm SQ↓.a"
    by simp
  ultimately have "SourceTerm SP↓.a"
      using reflection reflection_of_barbs_in_barbed_encoding[where Rel="Rel"]
    by blast
  thus "SP↓<SWB>a"
    by simp
qed

lemma (in encoding_wrt_barbs) indRelRSTPO_impl_SRel_and_TRel_reflect_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes reflection: "rel_reflects_barbs (indRelRSTPO SRel TRel) (STCalWB SWB TWB)"
  shows "rel_reflects_barbs SRel SWB"
    and "rel_reflects_barbs TRel TWB"
proof -
  show "rel_reflects_barbs SRel SWB"
      using reflection rel_with_source_impl_SRel_reflects_barbs[where
                        Rel="indRelRSTPO SRel TRel" and SRel="SRel"]
    by (simp add: indRelRSTPO.source)
next
  show "rel_reflects_barbs TRel TWB"
      using reflection rel_with_target_impl_TRel_reflects_barbs[where
                        Rel="indRelRSTPO SRel TRel" and TRel="TRel"]
    by (simp add: indRelRSTPO.target)
qed

lemma (in encoding_wrt_barbs) indRelLSTPO_impl_SRel_and_TRel_reflect_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes reflection: "rel_reflects_barbs (indRelLSTPO SRel TRel) (STCalWB SWB TWB)"
  shows "rel_reflects_barbs SRel SWB"
    and "rel_reflects_barbs TRel TWB"
proof -
  show "rel_reflects_barbs SRel SWB"
      using reflection rel_with_source_impl_SRel_reflects_barbs[where
                        Rel="indRelLSTPO SRel TRel" and SRel="SRel"]
    by (simp add: indRelLSTPO.source)
next
  show "rel_reflects_barbs TRel TWB"
      using reflection rel_with_target_impl_TRel_reflects_barbs[where
                        Rel="indRelLSTPO SRel TRel" and TRel="TRel"]
    by (simp add: indRelLSTPO.target)
qed

lemma (in encoding_wrt_barbs) indRelSTEQ_impl_SRel_and_TRel_reflect_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes reflection: "rel_reflects_barbs (indRelSTEQ SRel TRel) (STCalWB SWB TWB)"
  shows "rel_reflects_barbs SRel SWB"
    and "rel_reflects_barbs TRel TWB"
proof -
  show "rel_reflects_barbs SRel SWB"
      using reflection rel_with_source_impl_SRel_reflects_barbs[where
                        Rel="indRelSTEQ SRel TRel" and SRel="SRel"]
    by (simp add: indRelSTEQ.source)
next
  show "rel_reflects_barbs TRel TWB"
      using reflection rel_with_target_impl_TRel_reflects_barbs[where
                        Rel="indRelSTEQ SRel TRel" and TRel="TRel"]
    by (simp add: indRelSTEQ.target)
qed

lemma (in encoding_wrt_barbs) rel_with_source_impl_SRel_weakly_reflects_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and Rel  :: "(('procS, 'procT) Proc × ('procS, 'procT) Proc) set"
  assumes reflection: "rel_weakly_reflects_barbs Rel (STCalWB SWB TWB)"
      and sourceInRel:  "S1 S2. (S1, S2)  SRel  (SourceTerm S1, SourceTerm S2)  Rel"
  shows "rel_weakly_reflects_barbs SRel SWB"
proof clarify
  fix SP SQ a SQ'
  assume "(SP, SQ)  SRel"
  with sourceInRel have "(SourceTerm SP, SourceTerm SQ)  Rel"
    by blast
  moreover assume "SQ (Calculus SWB)* SQ'" and "SQ'↓<SWB>a"
  hence "SourceTerm SQ⇓.a"
    by blast
  ultimately have "SourceTerm SP⇓.a"
      using reflection weak_reflection_of_barbs_in_barbed_encoding[where Rel="Rel"]
    by blast
  thus "SP⇓<SWB>a"
    by simp
qed

lemma (in encoding_wrt_barbs) indRelRSTPO_impl_SRel_and_TRel_weakly_reflect_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes reflection: "rel_weakly_reflects_barbs (indRelRSTPO SRel TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_reflects_barbs SRel SWB"
    and "rel_weakly_reflects_barbs TRel TWB"
proof -
  show "rel_weakly_reflects_barbs SRel SWB"
      using reflection rel_with_source_impl_SRel_weakly_reflects_barbs[where
                        Rel="indRelRSTPO SRel TRel" and SRel="SRel"]
    by (simp add: indRelRSTPO.source)
next
  show "rel_weakly_reflects_barbs TRel TWB"
      using reflection rel_with_target_impl_TRel_weakly_reflects_barbs[where
                          Rel="indRelRSTPO SRel TRel" and TRel="TRel"]
    by (simp add: indRelRSTPO.target)
qed

lemma (in encoding_wrt_barbs) indRelLSTPO_impl_SRel_and_TRel_weakly_reflect_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes reflection: "rel_weakly_reflects_barbs (indRelLSTPO SRel TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_reflects_barbs SRel SWB"
    and "rel_weakly_reflects_barbs TRel TWB"
proof -
  show "rel_weakly_reflects_barbs SRel SWB"
      using reflection rel_with_source_impl_SRel_weakly_reflects_barbs[where
                        Rel="indRelLSTPO SRel TRel" and SRel="SRel"]
    by (simp add: indRelLSTPO.source)
next
  show "rel_weakly_reflects_barbs TRel TWB"
      using reflection rel_with_target_impl_TRel_weakly_reflects_barbs[where
                        Rel="indRelLSTPO SRel TRel" and TRel="TRel"]
    by (simp add: indRelLSTPO.target)
qed

lemma (in encoding_wrt_barbs) indRelSTEQ_impl_SRel_and_TRel_weakly_reflect_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes reflection: "rel_weakly_reflects_barbs (indRelSTEQ SRel TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_reflects_barbs SRel SWB"
    and "rel_weakly_reflects_barbs TRel TWB"
proof -
  show "rel_weakly_reflects_barbs SRel SWB"
      using reflection rel_with_source_impl_SRel_weakly_reflects_barbs[where
                        Rel="indRelSTEQ SRel TRel" and SRel="SRel"]
    by (simp add: indRelSTEQ.source)
next
  show "rel_weakly_reflects_barbs TRel TWB"
      using reflection rel_with_target_impl_TRel_weakly_reflects_barbs[where
                        Rel="indRelSTEQ SRel TRel" and TRel="TRel"]
    by (simp add: indRelSTEQ.target)
qed

text ‹If indRelRSTPO, indRelLSTPO, or indRelSTPO respects barbs then so do the corresponding
        source term and target term relations.›

lemma (in encoding_wrt_barbs) indRelRSTPO_impl_SRel_and_TRel_respect_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes respection: "rel_respects_barbs (indRelRSTPO SRel TRel) (STCalWB SWB TWB)"
  shows "rel_respects_barbs SRel SWB"
    and "rel_respects_barbs TRel TWB"
proof -
  show "rel_respects_barbs SRel SWB"
      using respection
            indRelRSTPO_impl_SRel_and_TRel_preserve_barbs(1)[where SRel="SRel" and TRel="TRel"]
            indRelRSTPO_impl_SRel_and_TRel_reflect_barbs(1)[where SRel="SRel" and TRel="TRel"]
    by blast
next
  show "rel_respects_barbs TRel TWB"
      using respection
            indRelRSTPO_impl_SRel_and_TRel_preserve_barbs(2)[where SRel="SRel" and TRel="TRel"]
            indRelRSTPO_impl_SRel_and_TRel_reflect_barbs(2)[where SRel="SRel" and TRel="TRel"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelLSTPO_impl_SRel_and_TRel_respect_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes respection: "rel_respects_barbs (indRelLSTPO SRel TRel) (STCalWB SWB TWB)"
  shows "rel_respects_barbs SRel SWB"
    and "rel_respects_barbs TRel TWB"
proof -
  show "rel_respects_barbs SRel SWB"
      using respection
            indRelLSTPO_impl_SRel_and_TRel_preserve_barbs(1)[where SRel="SRel" and TRel="TRel"]
            indRelLSTPO_impl_SRel_and_TRel_reflect_barbs(1)[where SRel="SRel" and TRel="TRel"]
    by blast
next
  show "rel_respects_barbs TRel TWB"
      using respection
            indRelLSTPO_impl_SRel_and_TRel_preserve_barbs(2)[where SRel="SRel" and TRel="TRel"]
            indRelLSTPO_impl_SRel_and_TRel_reflect_barbs(2)[where SRel="SRel" and TRel="TRel"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelSTEQ_impl_SRel_and_TRel_respect_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes respection: "rel_respects_barbs (indRelSTEQ SRel TRel) (STCalWB SWB TWB)"
  shows "rel_respects_barbs SRel SWB"
    and "rel_respects_barbs TRel TWB"
proof -
  show "rel_respects_barbs SRel SWB"
      using respection
            indRelSTEQ_impl_SRel_and_TRel_preserve_barbs(1)[where SRel="SRel" and TRel="TRel"]
            indRelSTEQ_impl_SRel_and_TRel_reflect_barbs(1)[where SRel="SRel" and TRel="TRel"]
    by blast
next
  show "rel_respects_barbs TRel TWB"
      using respection
            indRelSTEQ_impl_SRel_and_TRel_preserve_barbs(2)[where SRel="SRel" and TRel="TRel"]
            indRelSTEQ_impl_SRel_and_TRel_reflect_barbs(2)[where SRel="SRel" and TRel="TRel"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelRSTPO_impl_SRel_and_TRel_weakly_respect_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes respection: "rel_weakly_respects_barbs (indRelRSTPO SRel TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_respects_barbs SRel SWB"
    and "rel_weakly_respects_barbs TRel TWB"
proof -
  show "rel_weakly_respects_barbs SRel SWB"
      using respection indRelRSTPO_impl_SRel_and_TRel_weakly_preserve_barbs(1)[where SRel="SRel"
                        and TRel="TRel"]
            indRelRSTPO_impl_SRel_and_TRel_weakly_reflect_barbs(1)[where SRel="SRel"
              and TRel="TRel"]
    by blast
next
  show "rel_weakly_respects_barbs TRel TWB"
      using respection indRelRSTPO_impl_SRel_and_TRel_weakly_preserve_barbs(2)[where SRel="SRel"
                        and TRel="TRel"]
            indRelRSTPO_impl_SRel_and_TRel_weakly_reflect_barbs(2)[where SRel="SRel"
              and TRel="TRel"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelLSTPO_impl_SRel_and_TRel_weakly_respect_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes respection: "rel_weakly_respects_barbs (indRelLSTPO SRel TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_respects_barbs SRel SWB"
    and "rel_weakly_respects_barbs TRel TWB"
proof -
  show "rel_weakly_respects_barbs SRel SWB"
      using respection indRelLSTPO_impl_SRel_and_TRel_weakly_preserve_barbs(1)[where SRel="SRel"
                        and TRel="TRel"]
            indRelLSTPO_impl_SRel_and_TRel_weakly_reflect_barbs(1)[where SRel="SRel"
              and TRel="TRel"]
    by blast
next
  show "rel_weakly_respects_barbs TRel TWB"
      using respection indRelLSTPO_impl_SRel_and_TRel_weakly_preserve_barbs(2)[where SRel="SRel"
                        and TRel="TRel"]
            indRelLSTPO_impl_SRel_and_TRel_weakly_reflect_barbs(2)[where SRel="SRel"
              and TRel="TRel"]
    by blast
qed

lemma (in encoding_wrt_barbs) indRelSTEQ_impl_SRel_and_TRel_weakly_respect_barbs:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes respection: "rel_weakly_respects_barbs (indRelSTEQ SRel TRel) (STCalWB SWB TWB)"
  shows "rel_weakly_respects_barbs SRel SWB"
    and "rel_weakly_respects_barbs TRel TWB"
proof -
  show "rel_weakly_respects_barbs SRel SWB"
      using respection indRelSTEQ_impl_SRel_and_TRel_weakly_preserve_barbs(1)[where SRel="SRel"
                        and TRel="TRel"]
            indRelSTEQ_impl_SRel_and_TRel_weakly_reflect_barbs(1)[where SRel="SRel"
              and TRel="TRel"]
    by blast
next
  show "rel_weakly_respects_barbs TRel TWB"
      using respection indRelSTEQ_impl_SRel_and_TRel_weakly_preserve_barbs(2)[where SRel="SRel"
                        and TRel="TRel"]
            indRelSTEQ_impl_SRel_and_TRel_weakly_reflect_barbs(2)[where SRel="SRel"
              and TRel="TRel"]
    by blast
qed

text ‹If TRel is reflexive then ind relRTPO is a subrelation of indRelTEQ. If SRel is reflexive
        then indRelRTPO is a subrelation of indRelRTPO. Moreover, indRelRSTPO is a subrelation of
        indRelSTEQ.›

lemma (in encoding) indRelRTPO_to_indRelTEQ:
  fixes TRel :: "('procT × 'procT) set"
    and P Q  :: "('procS, 'procT) Proc"
  assumes rel:   "P ≲⟦⋅⟧RT<TRel> Q"
      and reflT: "refl TRel"
  shows "P ∼⟦⋅⟧T<TRel> Q"
    using rel
proof induct
  case (encR S)
  show "SourceTerm S ∼⟦⋅⟧T<TRel> TargetTerm (S)"
    by (rule indRelTEQ.encR)
next
  case (source S)
  from reflT show "SourceTerm S ∼⟦⋅⟧T<TRel> SourceTerm S"
      using indRelTEQ_refl[of TRel]
      unfolding refl_on_def
    by simp
next
  case (target T1 T2)
  assume "(T1, T2)  TRel"
  thus "TargetTerm T1 ∼⟦⋅⟧T<TRel> TargetTerm T2"
    by (rule indRelTEQ.target)
next
  case (trans TP TQ TR)
  assume "TP ∼⟦⋅⟧T<TRel> TQ" and "TQ ∼⟦⋅⟧T<TRel> TR"
  thus "TP ∼⟦⋅⟧T<TRel> TR"
    by (rule indRelTEQ.trans)
qed

lemma (in encoding) indRelRTPO_to_indRelRSTPO:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
    and P Q  :: "('procS, 'procT) Proc"
  assumes rel:   "P ≲⟦⋅⟧RT<TRel> Q"
      and reflS: "refl SRel"
  shows "P ≲⟦⋅⟧R<SRel,TRel> Q"
    using rel
proof induct
  case (encR S)
  show "SourceTerm S ≲⟦⋅⟧R<SRel,TRel> TargetTerm (S)"
    by (rule indRelRSTPO.encR)
next
  case (source S)
  from reflS show "SourceTerm S ≲⟦⋅⟧R<SRel,TRel> SourceTerm S"
      unfolding refl_on_def
    by (simp add: indRelRSTPO.source)
next
  case (target T1 T2)
  assume "(T1, T2)  TRel"
  thus "TargetTerm T1 ≲⟦⋅⟧R<SRel,TRel> TargetTerm T2"
    by (rule indRelRSTPO.target)
next
  case (trans P Q R)
  assume "P ≲⟦⋅⟧R<SRel,TRel> Q" and "Q ≲⟦⋅⟧R<SRel,TRel> R"
  thus "P ≲⟦⋅⟧R<SRel,TRel> R"
    by (rule indRelRSTPO.trans)
qed

lemma (in encoding) indRelRSTPO_to_indRelSTEQ:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
    and P Q  :: "('procS, 'procT) Proc"
  assumes rel: "P ≲⟦⋅⟧R<SRel,TRel> Q"
  shows "P ∼⟦⋅⟧<SRel,TRel> Q"
    using rel
proof induct
  case (encR S)
  show "SourceTerm S ∼⟦⋅⟧<SRel,TRel> TargetTerm (S)"
    by (rule indRelSTEQ.encR)
next
  case (source S1 S2)
  assume "(S1, S2)  SRel"
  thus "SourceTerm S1 ∼⟦⋅⟧<SRel,TRel> SourceTerm S2"
    by (rule indRelSTEQ.source)
next
  case (target T1 T2)
  assume "(T1, T2)  TRel"
  thus "TargetTerm T1 ∼⟦⋅⟧<SRel,TRel> TargetTerm T2"
    by (rule indRelSTEQ.target)
next
  case (trans P Q R)
  assume "P ∼⟦⋅⟧<SRel,TRel> Q" and "Q ∼⟦⋅⟧<SRel,TRel> R"
  thus "P ∼⟦⋅⟧<SRel,TRel> R"
    by (rule indRelSTEQ.trans)
qed

text ‹If indRelRTPO is a bisimulation and SRel is a reflexive bisimulation then also indRelRSTPO
        is a bisimulation.›

lemma (in encoding) indRelRTPO_weak_reduction_bisimulation_impl_indRelRSTPO_bisimulation:
  fixes SRel :: "('procS × 'procS) set"
    and TRel :: "('procT × 'procT) set"
  assumes bisimT: "weak_reduction_bisimulation (indRelRTPO TRel) (STCal Source Target)"
      and bisimS: "weak_reduction_bisimulation SRel Source"
      and reflS:  "refl SRel"
  shows "weak_reduction_bisimulation (indRelRSTPO SRel TRel) (STCal Source Target)"
proof auto
  fix P Q P'
  assume "P ≲⟦⋅⟧R<SRel,TRel> Q" and "P (STCal Source Target)* P'"
  thus "Q'. Q (STCal Source Target)* Q'  P' ≲⟦⋅⟧R<SRel,TRel> Q'"
  proof (induct arbitrary: P')
    case (encR S)
    have "SourceTerm S ≲⟦⋅⟧RT<TRel> TargetTerm (S)"
      by (rule indRelRTPO.encR)
    moreover assume "SourceTerm S (STCal Source Target)* P'"
    ultimately obtain Q' where A1: "TargetTerm (S) (STCal Source Target)* Q'"
                           and A2: "P' ≲⟦⋅⟧RT<TRel> Q'"
        using bisimT
      by blast
    from reflS A2 have "P' ≲⟦⋅⟧R<SRel,TRel> Q'"
      by (simp add: indRelRTPO_to_indRelRSTPO)
    with A1 show "Q'. TargetTerm (S) (STCal Source Target)* Q'  P' ≲⟦⋅⟧R<SRel,TRel> Q'"
      by blast
  next
    case (source S1 S2)
    assume "SourceTerm S1 (STCal Source Target)* P'"
    from this obtain S1' where B1: "S1' ∈S P'" and B2: "S1 Source* S1'"
      by (auto simp add: STCal_steps(1))
    assume "(S1, S2)  SRel"
    with B2 bisimS obtain S2' where B3: "S2 Source* S2'" and B4: "(S1', S2')  SRel"
      by blast
    from B3 have "SourceTerm S2 (STCal Source Target)* (SourceTerm S2')"
      by (simp add: STCal_steps(1))
    moreover from B1 B4 have "P' ≲⟦⋅⟧R<SRel,TRel> SourceTerm S2'"
      by (simp add: indRelRSTPO.source)
    ultimately show "Q'. SourceTerm S2 (STCal Source Target)* Q'  P' ≲⟦⋅⟧R<SRel,TRel> Q'"
      by blast
  next
    case (target T1 T2)
    assume "(T1, T2)  TRel"
    hence "TargetTerm T1 ≲⟦⋅⟧RT<TRel> TargetTerm T2"
      by (rule indRelRTPO.target)
    moreover assume "TargetTerm T1 (STCal Source Target)* P'"
    ultimately obtain Q' where C1: "TargetTerm T2 (STCal Source Target)* Q'"
                           and C2: "P' ≲⟦⋅⟧RT<TRel> Q'"
        using bisimT
      by blast
    from reflS C2 have "P' ≲⟦⋅⟧R<SRel,TRel> Q'"
      by (simp add: indRelRTPO_to_indRelRSTPO)
    with C1 show "Q'. TargetTerm T2 (STCal Source Target)* Q'  P' ≲⟦⋅⟧R<SRel,TRel> Q'"
      by blast
  next
    case (trans P Q R)
    assume "P (STCal Source Target)* P'"
       and "P'. P (STCal Source Target)* P'
             Q'. Q (STCal Source Target)* Q'  P' ≲⟦⋅⟧R<SRel,TRel> Q'"
    from this obtain Q' where D1: "Q (STCal Source Target)* Q'" and D2: "P' ≲⟦⋅⟧R<SRel,TRel> Q'"
      by blast
    assume "Q'. Q (STCal Source Target)* Q'
             R'. R (STCal Source Target)* R'  Q' ≲⟦⋅⟧R<SRel,TRel> R'"
    with D1 obtain R' where D3: "R (STCal Source Target)* R'" and D4: "Q' ≲⟦⋅⟧R<SRel,TRel> R'"
      by blast
    from D2 D4 have "P' ≲⟦⋅⟧R<SRel,TRel> R'"
      by (rule indRelRSTPO.trans)
    with D3 show "R'. R (STCal Source Target)* R'  P' ≲⟦⋅⟧R<SRel,TRel> R'"
      by blast
  qed
next
  fix P Q Q'
  assume "P ≲⟦⋅⟧R<SRel,TRel> Q" and "Q (STCal Source Target)* Q'"
  thus "P'. P (STCal Source Target)* P'  P' ≲⟦⋅⟧R<SRel,TRel> Q'"
  proof (induct arbitrary: Q')
    case (encR S)
    have "SourceTerm S ≲⟦⋅⟧RT<TRel> TargetTerm (S)"
      by (rule indRelRTPO.encR)
    moreover assume "TargetTerm (S) (STCal Source Target)* Q'"
    ultimately obtain P' where E1: "SourceTerm S (STCal Source Target)* P'"
                           and E2: "P' ≲⟦⋅⟧RT<TRel> Q'"
        using bisimT
      by blast
    from reflS E2 have "P' ≲⟦⋅⟧R<SRel,TRel> Q'"
      by (simp add: indRelRTPO_to_indRelRSTPO)
    with E1 show "P'. SourceTerm S (STCal Source Target)* P'  P' ≲⟦⋅⟧R<SRel,TRel> Q'"
      by blast
  next
    case (source S1 S2)
    assume "SourceTerm S2 (STCal Source Target)* Q'"
    from this obtain S2' where F1: "S2' ∈S Q'" and F2: "S2 Source* S2'"
      by (auto simp add: STCal_steps(1))
    assume "(S1, S2)  SRel"
    with F2 bisimS obtain S1' where F3: "S1 Source* S1'" and F4: "(S1', S2')  SRel"
      by blast
    from F3 have "SourceTerm S1 (STCal Source Target)* (SourceTerm S1')"
      by (simp add: STCal_steps(1))
    moreover from F1 F4 have "SourceTerm S1' ≲⟦⋅⟧R<SRel,TRel> Q'"
      by (simp add: indRelRSTPO.source)
    ultimately show "P'. SourceTerm S1 (STCal Source Target)* P'  P' ≲⟦⋅⟧R<SRel,TRel> Q'"
      by blast
  next
    case (target T1 T2)
    assume "(T1, T2)  TRel"
    hence "TargetTerm T1 ≲⟦⋅⟧RT<TRel> TargetTerm T2"
      by (rule indRelRTPO.target)
    moreover assume "TargetTerm T2 (STCal Source Target)* Q'"
    ultimately obtain P' where G1: "TargetTerm T1 (STCal Source Target)* P'"
                           and G2: "P' ≲⟦⋅⟧RT<TRel> Q'"
        using bisimT
      by blast
    from reflS G2 have "P' ≲⟦⋅⟧R<SRel,TRel> Q'"
      by (simp add: indRelRTPO_to_indRelRSTPO)
    with G1 show "P'. TargetTerm T1 (STCal Source Target)* P'  P' ≲⟦⋅⟧R<SRel,TRel> Q'"
      by blast
  next
    case (trans P Q R R')
    assume "R (STCal Source Target)* R'"
       and "R'. R (STCal Source Target)* R'
             Q'. Q (STCal Source Target)* Q'  Q' ≲⟦⋅⟧R<SRel,TRel> R'"
    from this obtain Q' where H1: "Q (STCal Source Target)* Q'" and H2: "Q' ≲⟦⋅⟧R<SRel,TRel> R'"
      by blast
    assume "Q'. Q (STCal Source Target)* Q'
             P'. P (STCal Source Target)* P'  P' ≲⟦⋅⟧R<SRel,TRel> Q'"
    with H1 obtain P' where H3: "P (STCal Source Target)* P'" and H4: "P' ≲⟦⋅⟧R<SRel,TRel> Q'"
      by blast
    from H4 H2 have "P' ≲⟦⋅⟧R<SRel,TRel> R'"
      by (rule indRelRSTPO.trans)
    with H3 show "P'. P (STCal Source Target)* P'  P' ≲⟦⋅⟧R<SRel,TRel> R'"
      by blast
  qed
qed

end