Theory Hygge_Theory

chapter‹Main Theorems›

theory Hygge_Theory
imports
  Corecursive_Prop

begin
text ‹
  Using the properties we have shown about the interpretation of configurations 
  and the stepwise unfolding of the denotational semantics, we can now prove 
  several important results about the construction of runs from a specification.
›
section ‹Initial configuration›

text ‹
  The denotational semantics of a specification @{term Ψ} is the interpretation 
  at the first instant of a configuration which has @{term Ψ} as its present.
  This means that we can start to build a run that satisfies a specification by 
  starting from this configuration.
›
theorem solve_start:
  shows ⟦⟦ Ψ ⟧⟧TESL =  [], 0  Ψ  [] config
  proof -
    have ⟦⟦ Ψ ⟧⟧TESL = ⟦⟦ Ψ ⟧⟧TESL⇗≥ 0⇖›
    by (simp add: TESL_interpretation_stepwise_zero')
    moreover have  [], 0  Ψ  [] config =
                      ⟦⟦ [] ⟧⟧prim  ⟦⟦ Ψ ⟧⟧TESL⇗≥ 0 ⟦⟦ [] ⟧⟧TESL⇗≥ Suc 0⇖›
    by simp
    ultimately show ?thesis by auto
  qed

section ‹Soundness›
text ‹
  The interpretation of a configuration @{term 𝒮2} that is a refinement of a
  configuration @{term 𝒮1} is contained in the interpretation of @{term 𝒮1}.
  This means that by making successive choices in building the instants of a run,
  we preserve the soundness of the constructed run with regard to the original 
  specification.
›
lemma sound_reduction:
  assumes (Γ1, n1  Ψ1  Φ1)    (Γ2, n2  Ψ2  Φ2)
  shows ⟦⟦ Γ1 ⟧⟧prim  ⟦⟦ Ψ1 ⟧⟧TESL⇗≥ n1 ⟦⟦ Φ1 ⟧⟧TESL⇗≥ Suc n1  ⟦⟦ Γ2 ⟧⟧prim  ⟦⟦ Ψ2 ⟧⟧TESL⇗≥ n2 ⟦⟦ Φ2 ⟧⟧TESL⇗≥ Suc n2⇖› (is ?P)
proof -
  from assms consider
    (a) (Γ1, n1  Ψ1  Φ1)  i  (Γ2, n2  Ψ2  Φ2)
  | (b) (Γ1, n1  Ψ1  Φ1)  e  (Γ2, n2  Ψ2  Φ2)
    using operational_semantics_step.simps by blast
  thus ?thesis
  proof (cases)
    case a
      thus ?thesis by (simp add: operational_semantics_intro.simps)
  next
    case b thus ?thesis
    proof (rule operational_semantics_elim.cases)
      fix  Γ n K1 τ K2 Ψ Φ
      assume (Γ1, n1  Ψ1  Φ1) = (Γ, n  (K1 sporadic τ on K2) # Ψ  Φ)
      and (Γ2, n2  Ψ2  Φ2) = (Γ, n  Ψ  ((K1 sporadic τ on K2) # Φ))
      thus ?P using HeronConf_interp_stepwise_sporadicon_cases
                    HeronConf_interpretation.simps by blast
    next
      fix  Γ n K1 τ K2 Ψ Φ
      assume (Γ1, n1  Ψ1  Φ1) = (Γ, n  (K1 sporadic τ on K2) # Ψ  Φ)
      and (Γ2, n2  Ψ2  Φ2) = (((K1  n) # (K2  n @ τ) # Γ), n  Ψ  Φ)
      thus ?P using HeronConf_interp_stepwise_sporadicon_cases
                    HeronConf_interpretation.simps by blast
    next
      fix Γ n K1 K2 R Ψ Φ
      assume (Γ1, n1  Ψ1  Φ1) = (Γ, n  (time-relation K1, K2  R) # Ψ  Φ)
      and (Γ2, n2  Ψ2  Φ2) = (((τvar (K1, n), τvar (K2, n)  R) # Γ), n
                                       Ψ  ((time-relation K1, K2  R) # Φ))
      thus ?P using HeronConf_interp_stepwise_tagrel_cases
                    HeronConf_interpretation.simps by blast
    next
      fix Γ n K1 K2 Ψ Φ
      assume (Γ1, n1  Ψ1  Φ1) = (Γ, n  (K1 implies K2) # Ψ  Φ)
      and (Γ2, n2  Ψ2  Φ2) = (((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 implies K2) # Φ))
      thus ?P using HeronConf_interp_stepwise_implies_cases
                    HeronConf_interpretation.simps by blast
    next
      fix Γ n K1 K2 Ψ Φ
      assume (Γ1, n1  Ψ1  Φ1) = (Γ, n  ((K1 implies K2) # Ψ)  Φ)
      and (Γ2, n2  Ψ2  Φ2) = (((K1  n) # (K2  n) # Γ), n
                                    Ψ  ((K1 implies K2) # Φ))
      thus ?P using HeronConf_interp_stepwise_implies_cases
                    HeronConf_interpretation.simps by blast
    next
      fix Γ n K1 K2 Ψ Φ
      assume (Γ1, n1  Ψ1  Φ1) = (Γ, n  ((K1 implies not K2) # Ψ)  Φ)
      and (Γ2, n2  Ψ2  Φ2) = (((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 implies not K2) # Φ))
      thus ?P using HeronConf_interp_stepwise_implies_not_cases
                    HeronConf_interpretation.simps by blast
    next
      fix Γ n K1 K2 Ψ Φ
      assume (Γ1, n1  Ψ1  Φ1) = (Γ, n  ((K1 implies not K2) # Ψ)  Φ)
      and (Γ2, n2  Ψ2  Φ2) = (((K1  n) # (K2 ¬⇑ n) # Γ), n
                                    Ψ  ((K1 implies not K2) # Φ))
      thus ?P using HeronConf_interp_stepwise_implies_not_cases
                    HeronConf_interpretation.simps by blast
    next
      fix Γ n K1 δτ K2 K3 Ψ Φ
      assume (Γ1, n1  Ψ1  Φ1) =
                (Γ, n  ((K1 time-delayed by δτ on K2 implies K3) # Ψ)  Φ)
      and (Γ2, n2  Ψ2  Φ2) =
            (((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 time-delayed by δτ on K2 implies K3) # Φ))
      thus ?P using HeronConf_interp_stepwise_timedelayed_cases
                    HeronConf_interpretation.simps by blast
    next
      fix Γ n K1 δτ K2 K3 Ψ Φ
      assume (Γ1, n1  Ψ1  Φ1) =
               (Γ, n  ((K1 time-delayed by δτ on K2 implies K3) # Ψ)  Φ)
      and (Γ2, n2  Ψ2  Φ2)
            = (((K1  n) # (K2 @ n  δτ  K3) # Γ), n
                 Ψ  ((K1 time-delayed by δτ on K2 implies K3) # Φ))
      thus ?P using HeronConf_interp_stepwise_timedelayed_cases
                    HeronConf_interpretation.simps by blast
    next
      fix Γ n K1 K2 Ψ Φ
      assume (Γ1, n1  Ψ1  Φ1) = (Γ, n  ((K1 weakly precedes K2) # Ψ)  Φ)
      and (Γ2, n2  Ψ2  Φ2) = (((# K2 n, # K1 n  (λ(x, y). x  y)) # Γ), n
                                   Ψ  ((K1 weakly precedes K2) # Φ))
      thus ?P using HeronConf_interp_stepwise_weakly_precedes_cases
                    HeronConf_interpretation.simps by blast
    next
      fix Γ n K1 K2 Ψ Φ
      assume (Γ1, n1  Ψ1  Φ1) = (Γ, n  ((K1 strictly precedes K2) # Ψ)  Φ)
      and (Γ2, n2  Ψ2  Φ2) = (((# K2 n, #< K1 n  (λ(x, y). x  y)) # Γ), n
                                   Ψ  ((K1 strictly precedes K2) # Φ))
      thus ?P using HeronConf_interp_stepwise_strictly_precedes_cases
                    HeronConf_interpretation.simps by blast
    next
      fix Γ n K1 K2 Ψ Φ
      assume (Γ1, n1  Ψ1  Φ1) = (Γ, n  ((K1 kills K2) # Ψ)  Φ)
      and (Γ2, n2  Ψ2  Φ2) = (((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 kills K2) # Φ))
      thus ?P using HeronConf_interp_stepwise_kills_cases
                    HeronConf_interpretation.simps by blast
    next
      fix Γ n K1 K2 Ψ Φ
      assume (Γ1, n1  Ψ1  Φ1) = (Γ, n  ((K1 kills K2) # Ψ)  Φ)
      and (Γ2, n2  Ψ2  Φ2) =
            (((K1  n) # (K2 ¬⇑  n) # Γ), n  Ψ  ((K1 kills K2) # Φ))
      thus ?P using HeronConf_interp_stepwise_kills_cases
                    HeronConf_interpretation.simps by blast
    qed
  qed
qed

inductive_cases step_elim:𝒮1  𝒮2

lemma sound_reduction':
  assumes 𝒮1  𝒮2
  shows  𝒮1 config   𝒮2 config
proof -
  have s1 s2. ( s2 config   s1 config)  ¬(s1  s2)
    using sound_reduction by fastforce
  thus ?thesis using assms by blast
qed

lemma sound_reduction_generalized:
  assumes 𝒮1 ↪⇗k𝒮2
    shows  𝒮1 config   𝒮2 config
proof -
  from assms show ?thesis
  proof (induction k arbitrary: 𝒮2)
    case 0
      hence *: 𝒮1 ↪⇗0𝒮2  𝒮1 = 𝒮2 by auto
      moreover have 𝒮1 = 𝒮2 using * "0.prems" by linarith
      ultimately show ?case by auto
  next
    case (Suc k)
      thus ?case
      proof -
        fix k :: nat
        assume ff: 𝒮1 ↪⇗Suc k𝒮2
        assume hi: 𝒮2. 𝒮1 ↪⇗k𝒮2   𝒮2 config   𝒮1 config
        obtain 𝒮n where red_decomp: (𝒮1 ↪⇗k𝒮n)  (𝒮n  𝒮2) using ff by auto
        hence  𝒮1 config   𝒮n config using hi by simp
        also have  𝒮n config   𝒮2 config by (simp add: red_decomp sound_reduction')
        ultimately show  𝒮1 config   𝒮2 config by simp
      qed
  qed
qed

text ‹
  From the initial configuration, a configuration @{term 𝒮} obtained after any 
  number @{term k} of reduction steps denotes runs from the 
  initial specification @{term Ψ}.
›
theorem soundness:
  assumes ([], 0  Ψ  []) ↪⇗k𝒮
    shows ⟦⟦ Ψ ⟧⟧TESL   𝒮 config
  using assms sound_reduction_generalized solve_start by blast

section ‹Completeness›

text ‹
  We will now show that any run that satisfies a specification can be derived
  from the initial configuration, at any number of steps.

  We start by proving that any run that is denoted by a configuration @{term 𝒮}
  is necessarily denoted by at least one of the configurations that can be reached
  from @{term 𝒮}.
›
lemma complete_direct_successors:
  shows  Γ, n  Ψ  Φ config  (X𝒞next (Γ, n  Ψ  Φ).  X config)
  proof (induct Ψ)
    case Nil
    show ?case
      using HeronConf_interp_stepwise_instant_cases operational_semantics_step.simps
            operational_semantics_intro.instant_i
      by fastforce
  next
    case (Cons ψ Ψ)  thus ?case
      proof (cases ψ)
        case (SporadicOn K1 τ K2) thus ?thesis 
          using HeronConf_interp_stepwise_sporadicon_cases
                                      [of Γ n K1 τ K2 Ψ Φ]
                Cnext_solve_sporadicon[of Γ n Ψ K1 τ K2 Φ] by blast
      next
        case (TagRelation K1 K2 R) thus ?thesis
          using HeronConf_interp_stepwise_tagrel_cases
                                  [of Γ n K1 K2 R Ψ Φ]
                Cnext_solve_tagrel[of K1 n K2 R Γ Ψ Φ] by blast
      next
        case (Implies K1 K2) thus ?thesis
          using HeronConf_interp_stepwise_implies_cases
                                   [of Γ n K1 K2 Ψ Φ]
                Cnext_solve_implies[of K1 n Γ Ψ K2 Φ] by blast
      next
        case (ImpliesNot K1 K2) thus ?thesis
          using HeronConf_interp_stepwise_implies_not_cases
                                       [of Γ n K1 K2 Ψ Φ]
                Cnext_solve_implies_not[of K1 n Γ Ψ K2 Φ] by blast
      next
        case (TimeDelayedBy Kmast τ Kmeas Kslave) thus ?thesis
          using HeronConf_interp_stepwise_timedelayed_cases
                          [of Γ n Kmast τ Kmeas Kslave Ψ Φ]
                Cnext_solve_timedelayed
                          [of Kmast n Γ Ψ τ Kmeas Kslave Φ] by blast
      next
        case (WeaklyPrecedes K1 K2) thus ?thesis
          using HeronConf_interp_stepwise_weakly_precedes_cases
                                           [of Γ n K1 K2 Ψ Φ]
                Cnext_solve_weakly_precedes[of K2 n K1 Γ Ψ  Φ]
          by blast
      next
        case (StrictlyPrecedes K1 K2)  thus ?thesis
          using HeronConf_interp_stepwise_strictly_precedes_cases
                                             [of Γ n K1 K2 Ψ Φ]
                Cnext_solve_strictly_precedes[of K2 n K1 Γ Ψ  Φ]
          by blast
      next
        case (Kills K1 K2) thus ?thesis
          using HeronConf_interp_stepwise_kills_cases[of Γ n K1 K2 Ψ Φ]
                Cnext_solve_kills[of K1 n Γ Ψ K2 Φ] by blast
      qed
  qed

lemma complete_direct_successors':
  shows  𝒮 config  (X𝒞next 𝒮.  X config)
proof -
  from HeronConf_interpretation.cases obtain Γ n Ψ Φ
    where 𝒮 = (Γ, n  Ψ  Φ) by blast
  with complete_direct_successors[of Γ n Ψ Φ] show ?thesis by simp
qed

text ‹
  Therefore, if a run belongs to a configuration, it necessarily belongs to a
  configuration derived from it.
›
lemma branch_existence:
  assumes ρ   𝒮1 config
  shows 𝒮2. (𝒮1  𝒮2)  (ρ   𝒮2 config)
proof -
  from assms complete_direct_successors' have ρ  (X𝒞next 𝒮1.  X config) by blast
  hence s𝒞next 𝒮1. ρ   s config by simp
  thus ?thesis by blast
qed

lemma branch_existence':
  assumes ρ   𝒮1 config
  shows 𝒮2. (𝒮1 ↪⇗k𝒮2)  (ρ   𝒮2 config)
proof (induct k)
  case 0
    thus ?case by (simp add: assms)
next
  case (Suc k)
    thus ?case
      using branch_existence relpowp_Suc_I[of k operational_semantics_step]
    by blast
qed

text‹
  Any run that belongs to the original specification @{term Ψ} has a corresponding 
  configuration @{term 𝒮} at any number @{term k} of reduction steps
  from the initial configuration. Therefore, any run that satisfies a specification
  can be derived from the initial configuration at any level of reduction.
›
theorem completeness:
  assumes ρ  ⟦⟦ Ψ ⟧⟧TESL
  shows 𝒮. (([], 0  Ψ  [])  ↪⇗k𝒮)
            ρ   𝒮 config
  using assms branch_existence' solve_start by blast

section ‹Progress›

text ‹
  Reduction steps do not guarantee that the construction of a run progresses in the
  sequence of instants. We need to show that it is always possible to reach the next 
  instant, and therefore any future instant, through a number of steps.
›
lemma instant_index_increase:
  assumes ρ   Γ, n  Ψ  Φ config
  shows   Γk Ψk Φk k. ((Γ, n  Ψ  Φ)  ↪⇗k(Γk, Suc n  Ψk  Φk))
                          ρ   Γk, Suc n  Ψk  Φk config
proof (insert assms, induct Ψ arbitrary: Γ Φ)
  case (Nil Γ Φ)
    then show ?case
    proof -
      have (Γ, n  []  Φ) ↪⇗1(Γ, Suc n  Φ  [])
        using instant_i intro_part by fastforce
      moreover have  Γ, n  []  Φ config =  Γ, Suc n  Φ  [] config
        by auto
      moreover have ρ   Γ, Suc n  Φ  [] config
        using assms Nil.prems calculation(2) by blast
      ultimately show ?thesis by blast
    qed
next
  case (Cons ψ Ψ)
    then show ?case
    proof (induct ψ)
      case (SporadicOn K1 τ K2)
        have branches:  Γ, n  ((K1 sporadic τ on K2) # Ψ)  Φ config
                      =  Γ, n  Ψ  ((K1 sporadic τ on K2) # Φ) config
                        ((K1  n) # (K2  n @ τ) # Γ), n  Ψ  Φ config
          using HeronConf_interp_stepwise_sporadicon_cases by simp
        have br1: ρ   Γ, n  Ψ  ((K1 sporadic τ on K2) # Φ) config
                     Γk Ψk Φk k.
                      ((Γ, n  ((K1 sporadic τ on K2) # Ψ)  Φ)
                          ↪⇗k(Γk, Suc n  Ψk  Φk))
                       ρ   Γk, Suc n  Ψk  Φk config
        proof -
          assume h1: ρ   Γ, n  Ψ  ((K1 sporadic τ on K2) # Φ) config
          hence Γk Ψk Φk k. ((Γ, n  Ψ  ((K1 sporadic τ on K2) # Φ))
                                      ↪⇗k(Γk, Suc n  Ψk  Φk))
                                  (ρ   Γk, Suc n  Ψk  Φk config)
            using h1 SporadicOn.prems by simp
          from this obtain Γk Ψk Φk k where
              fp:((Γ, n  Ψ  ((K1 sporadic τ on K2) # Φ))
                    ↪⇗k(Γk, Suc n  Ψk  Φk))
                 ρ   Γk, Suc n  Ψk  Φk config by blast
          have
            (Γ, n  ((K1 sporadic τ on K2) # Ψ)  Φ)
               (Γ, n  Ψ  ((K1 sporadic τ on K2) # Φ))
            by (simp add: elims_part sporadic_on_e1)
          with fp relpowp_Suc_I2 have
            ((Γ, n  ((K1 sporadic τ on K2) # Ψ)  Φ)
              ↪⇗Suc k(Γk, Suc n  Ψk  Φk)) by auto
          thus ?thesis using fp by blast
        qed
        have br2: ρ   ((K1  n) # (K2  n @ τ) # Γ), n  Ψ  Φ config
                   Γk Ψk Φk k. ((Γ, n  ((K1 sporadic τ on K2) # Ψ)  Φ)
                                        ↪⇗k(Γk, Suc n  Ψk  Φk))
                             ρ   Γk, Suc n  Ψk  Φk config
        proof -
          assume h2: ρ   ((K1  n) # (K2  n @ τ) # Γ), n  Ψ  Φ config
          hence Γk Ψk Φk k. ((((K1  n) # (K2  n @ τ) # Γ), n  Ψ  Φ)
                                    ↪⇗k(Γk, Suc n  Ψk  Φk))
                              ρ   Γk, Suc n  Ψk  Φk config
            using h2 SporadicOn.prems by simp

            from this obtain Γk Ψk Φk k
            where fp:((((K1  n) # (K2  n @ τ) # Γ), n  Ψ  Φ)
                            ↪⇗k(Γk, Suc n  Ψk  Φk))
              and rc:ρ   Γk, Suc n  Ψk  Φk config by blast
            have pc:(Γ, n  ((K1 sporadic τ on K2) # Ψ)  Φ)
               (((K1  n) # (K2  n @ τ) # Γ), n  Ψ  Φ)
            by (simp add: elims_part sporadic_on_e2)
            hence (Γ, n  (K1 sporadic τ on K2) # Ψ  Φ)
                    ↪⇗Suc k(Γk, Suc n  Ψk  Φk)
                using fp relpowp_Suc_I2 by auto
            with rc show ?thesis by blast
        qed
        from branches SporadicOn.prems(2) have
          ρ   Γ, n  Ψ  ((K1 sporadic τ on K2) # Φ) config
               ((K1  n) # (K2  n @ τ) # Γ), n  Ψ  Φ config
          by simp
        with br1 br2 show ?case by blast
  next
    case (TagRelation K1 K2 R)
      have branches:  Γ, n  ((time-relation K1, K2  R) # Ψ)  Φ config
          =  ((τvar(K1, n), τvar(K2, n)  R) # Γ), n
               Ψ  ((time-relation K1, K2  R) # Φ) config
        using HeronConf_interp_stepwise_tagrel_cases by simp
      thus ?case
      proof -
        have Γk Ψk Φk k.
            ((((τvar(K1, n), τvar(K2, n)  R) # Γ), n
                 Ψ  ((time-relation K1, K2  R) # Φ))
              ↪⇗k(Γk, Suc n  Ψk  Φk))  ρ   Γk, Suc n  Ψk  Φk config
          using TagRelation.prems by simp

        from this obtain Γk Ψk Φk k
          where fp:((((τvar(K1, n), τvar(K2, n)  R) # Γ), n
                         Ψ  ((time-relation K1, K2  R) # Φ))
                    ↪⇗k(Γk, Suc n  Ψk  Φk))
            and rc:ρ   Γk, Suc n  Ψk  Φk config by blast
        have pc:(Γ, n  ((time-relation K1, K2  R) # Ψ)  Φ)
             (((τvar (K1, n), τvar (K2, n)  R) # Γ), n
                   Ψ  ((time-relation K1, K2  R) # Φ))
          by (simp add: elims_part tagrel_e)
        hence (Γ, n  (time-relation K1, K2  R) # Ψ  Φ)
                ↪⇗Suc k(Γk, Suc n  Ψk  Φk)
          using fp relpowp_Suc_I2 by auto
        with rc show ?thesis by blast
      qed
  next
    case (Implies K1 K2)
      have branches:  Γ, n  ((K1 implies K2) # Ψ)  Φ config
          =  ((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 implies K2) # Φ) config
            ((K1  n) # (K2  n) # Γ), n  Ψ  ((K1 implies K2) # Φ) config
        using HeronConf_interp_stepwise_implies_cases by simp
      moreover have br1: ρ   ((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 implies K2) # Φ) config
                 Γk Ψk Φk k. ((Γ, n  ((K1 implies K2) # Ψ)  Φ)
                                    ↪⇗k(Γk, Suc n  Ψk  Φk))
                   ρ   Γk, Suc n  Ψk  Φk config
      proof -
        assume h1: ρ   ((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 implies K2) # Φ) config
        then have Γk Ψk Φk k.
                    ((((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 implies K2) # Φ))
                        ↪⇗k(Γk, Suc n  Ψk  Φk))
                   ρ   Γk, Suc n  Ψk  Φk config
          using h1 Implies.prems by simp
        from this obtain Γk Ψk Φk k where
          fp:((((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 implies K2) # Φ))
            ↪⇗k(Γk, Suc n  Ψk  Φk))
          and rc:ρ   Γk, Suc n  Ψk  Φk config by blast
        have pc:(Γ, n  (K1 implies K2) # Ψ  Φ)
                   (((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 implies K2) # Φ))
          by (simp add: elims_part implies_e1)
        hence (Γ, n  (K1 implies K2) # Ψ  Φ) ↪⇗Suc k(Γk, Suc n  Ψk  Φk)
          using fp relpowp_Suc_I2 by auto
        with rc show ?thesis by blast
      qed
      moreover have br2: ρ   ((K1  n) # (K2  n) # Γ), n
                                 Ψ  ((K1 implies K2) # Φ) config
                             Γk Ψk Φk k. ((Γ, n  ((K1 implies K2) # Ψ)  Φ)
                                              ↪⇗k(Γk, Suc n  Ψk  Φk))
                                   ρ   Γk, Suc n  Ψk  Φk config
      proof -
        assume h2: ρ   ((K1  n) # (K2  n) # Γ), n
                           Ψ  ((K1 implies K2) # Φ) config
        then have Γk Ψk Φk k. (
                        (((K1  n) # (K2  n) # Γ), n  Ψ  ((K1 implies K2) # Φ))
                          ↪⇗k(Γk, Suc n  Ψk  Φk)
                    )  ρ   Γk, Suc n  Ψk  Φk config
          using h2 Implies.prems by simp
        from this obtain Γk Ψk Φk k where
            fp:(((K1  n) # (K2  n) # Γ), n  Ψ  ((K1 implies K2) # Φ))
                ↪⇗k(Γk, Suc n  Ψk  Φk)
        and rc:ρ   Γk, Suc n  Ψk  Φk config by blast
        have (Γ, n  ((K1 implies K2) # Ψ)  Φ)
               (((K1  n) # (K2  n) # Γ), n  Ψ  ((K1 implies K2) # Φ))
          by (simp add: elims_part implies_e2)
        hence (Γ, n  ((K1 implies K2) # Ψ)  Φ) ↪⇗Suc k(Γk, Suc n  Ψk  Φk)
          using fp relpowp_Suc_I2 by auto
        with rc show ?thesis by blast
      qed
      ultimately show ?case using Implies.prems(2) by blast
  next
    case (ImpliesNot K1 K2)
      have branches:  Γ, n  ((K1 implies not K2) # Ψ)  Φ config
          =  ((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 implies not K2) # Φ) config
            ((K1  n) # (K2 ¬⇑ n) # Γ), n  Ψ  ((K1 implies not K2) # Φ) config
        using HeronConf_interp_stepwise_implies_not_cases by simp
      moreover have br1: ρ   ((K1 ¬⇑ n) # Γ), n
                             Ψ  ((K1 implies not K2) # Φ) config
                 Γk Ψk Φk k. ((Γ, n  ((K1 implies not K2) # Ψ)  Φ)
                                    ↪⇗k(Γk, Suc n  Ψk  Φk))
                   ρ   Γk, Suc n  Ψk  Φk config
      proof -
        assume h1: ρ   ((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 implies not K2) # Φ) config
        then have Γk Ψk Φk k.
                    ((((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 implies not K2) # Φ))
                      ↪⇗k(Γk, Suc n  Ψk  Φk))
                   ρ   Γk, Suc n  Ψk  Φk config
          using h1 ImpliesNot.prems by simp
        from this obtain Γk Ψk Φk k where
          fp:((((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 implies not K2) # Φ))
                ↪⇗k(Γk, Suc n  Ψk  Φk))
          and rc:ρ   Γk, Suc n  Ψk  Φk config by blast
        have pc:(Γ, n  (K1 implies not K2) # Ψ  Φ)
                   (((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 implies not K2) # Φ))
          by (simp add: elims_part implies_not_e1)
        hence (Γ, n  (K1 implies not K2) # Ψ  Φ) ↪⇗Suc k(Γk, Suc n  Ψk  Φk)
          using fp relpowp_Suc_I2 by auto
        with rc show ?thesis by blast
      qed
      moreover have br2: ρ   ((K1  n) # (K2 ¬⇑ n) # Γ), n
                             Ψ  ((K1 implies not K2) # Φ) config
                             Γk Ψk Φk k. ((Γ, n  ((K1 implies not K2) # Ψ)  Φ)
                                              ↪⇗k(Γk, Suc n  Ψk  Φk))
                                   ρ   Γk, Suc n  Ψk  Φk config
      proof -
        assume h2: ρ   ((K1  n) # (K2 ¬⇑ n) # Γ), n
                           Ψ  ((K1 implies not K2) # Φ) config
        then have Γk Ψk Φk k. (
                     (((K1  n) # (K2 ¬⇑ n) # Γ), n
                        Ψ  ((K1 implies not K2) # Φ)) ↪⇗k(Γk, Suc n  Ψk  Φk)
                    )  ρ   Γk, Suc n  Ψk  Φk config
          using h2 ImpliesNot.prems by simp
        from this obtain Γk Ψk Φk k where
            fp:(((K1  n) # (K2 ¬⇑ n) # Γ), n  Ψ  ((K1 implies not K2) # Φ))
                ↪⇗k(Γk, Suc n  Ψk  Φk)
        and rc:ρ   Γk, Suc n  Ψk  Φk config by blast
        have (Γ, n  ((K1 implies not K2) # Ψ)  Φ)
               (((K1  n) # (K2 ¬⇑ n) # Γ), n  Ψ  ((K1 implies not K2) # Φ))
          by (simp add: elims_part implies_not_e2)
        hence (Γ, n  ((K1 implies not K2) # Ψ)  Φ)
                ↪⇗Suc k(Γk, Suc n  Ψk  Φk)
          using fp relpowp_Suc_I2 by auto
        with rc show ?thesis by blast
      qed
      ultimately show ?case  using ImpliesNot.prems(2) by blast
  next
    case (TimeDelayedBy K1 δτ K2 K3)
      have branches:
         Γ, n  ((K1 time-delayed by δτ on K2 implies K3) # Ψ)  Φ config
          =  ((K1 ¬⇑ n) # Γ), n
               Ψ  ((K1 time-delayed by δτ on K2 implies K3) # Φ) config
            ((K1  n) # (K2 @ n  δτ  K3) # Γ), n
               Ψ  ((K1 time-delayed by δτ on K2 implies K3) # Φ) config
        using HeronConf_interp_stepwise_timedelayed_cases by simp
      moreover have br1:
        ρ   ((K1 ¬⇑ n) # Γ), n
               Ψ  ((K1 time-delayed by δτ on K2 implies K3) # Φ) config
           Γk Ψk Φk k.
            ((Γ, n  ((K1 time-delayed by δτ on K2 implies K3) # Ψ)  Φ)
                ↪⇗k(Γk, Suc n  Ψk  Φk))
             ρ   Γk, Suc n  Ψk  Φk config
      proof -
        assume h1: ρ   ((K1 ¬⇑ n) # Γ), n
                          Ψ  ((K1 time-delayed by δτ on K2 implies K3) # Φ) config
        then have Γk Ψk Φk k.
          ((((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 time-delayed by δτ on K2 implies K3) # Φ))
            ↪⇗k(Γk, Suc n  Ψk  Φk))
           ρ   Γk, Suc n  Ψk  Φk config
          using h1 TimeDelayedBy.prems by simp
        from this obtain Γk Ψk Φk k
          where fp:(((K1 ¬⇑ n) # Γ), n
                       Ψ  ((K1 time-delayed by δτ on K2 implies K3) # Φ))
                    ↪⇗k(Γk, Suc n  Ψk  Φk)
            and rc:ρ   Γk, Suc n  Ψk  Φk config by blast
        have (Γ, n  ((K1 time-delayed by δτ on K2 implies K3) # Ψ)  Φ)
               (((K1 ¬⇑ n) # Γ), n
                     Ψ  ((K1 time-delayed by δτ on K2 implies K3) # Φ))
          by (simp add: elims_part timedelayed_e1)
        hence (Γ, n  ((K1 time-delayed by δτ on K2 implies K3) # Ψ)  Φ)
                ↪⇗Suc k(Γk, Suc n  Ψk  Φk)
          using fp relpowp_Suc_I2 by auto
        with rc show ?thesis by blast
      qed
      moreover have br2:
        ρ   ((K1  n) # (K2 @ n  δτ  K3) # Γ), n
               Ψ  ((K1 time-delayed by δτ on K2 implies K3) # Φ) config
           Γk Ψk Φk k.
              ((Γ, n  ((K1 time-delayed by δτ on K2 implies K3) # Ψ)  Φ)
                ↪⇗k(Γk, Suc n  Ψk  Φk))
               ρ   Γk, Suc n  Ψk  Φk config
      proof -
        assume h2: ρ   ((K1  n) # (K2 @ n  δτ  K3) # Γ), n
                     Ψ  ((K1 time-delayed by δτ on K2 implies K3) # Φ) config
        then have Γk Ψk Φk k. ((((K1  n) # (K2 @ n  δτ  K3) # Γ), n
                                  Ψ  ((K1 time-delayed by δτ on K2 implies K3) # Φ))
                                 ↪⇗k(Γk, Suc n  Ψk  Φk))
                                 ρ   Γk, Suc n  Ψk  Φk config
          using h2 TimeDelayedBy.prems by simp
        from this obtain Γk Ψk Φk k
          where fp:(((K1  n) # (K2 @ n  δτ  K3) # Γ), n
                          Ψ  ((K1 time-delayed by δτ on K2 implies K3) # Φ))
                       ↪⇗k(Γk, Suc n  Ψk  Φk)
            and rc:ρ   Γk, Suc n  Ψk  Φk config by blast
        have (Γ, n  ((K1 time-delayed by δτ on K2 implies K3) # Ψ)  Φ)
               (((K1  n) # (K2 @ n  δτ  K3) # Γ), n
                   Ψ  ((K1 time-delayed by δτ on K2 implies K3) # Φ))
          by (simp add: elims_part timedelayed_e2)
        with fp relpowp_Suc_I2 have
          (Γ, n  ((K1 time-delayed by δτ on K2 implies K3) # Ψ)  Φ)
            ↪⇗Suc k(Γk, Suc n  Ψk  Φk)
          by auto
        with rc show ?thesis by blast
      qed
      ultimately show ?case using TimeDelayedBy.prems(2) by blast
  next
    case (WeaklyPrecedes K1 K2)
      have  Γ, n  ((K1 weakly precedes K2) # Ψ)  Φ config =
         ((# K2 n, # K1 n  (λ(x, y). x  y)) # Γ), n
             Ψ  ((K1 weakly precedes K2) # Φ) config
        using HeronConf_interp_stepwise_weakly_precedes_cases by simp
      moreover have ρ   ((# K2 n, # K1 n  (λ(x, y). x  y)) # Γ), n
                             Ψ  ((K1 weakly precedes K2) # Φ) config
             (Γk Ψk Φk k. ((Γ, n  ((K1 weakly precedes K2) # Ψ)  Φ)
                                  ↪⇗k(Γk, Suc n  Ψk  Φk))
                 (ρ   Γk, Suc n  Ψk  Φk config))
      proof -
        assume ρ   ((# K2 n, # K1 n  (λ(x, y). x  y)) # Γ), n
                         Ψ  ((K1 weakly precedes K2) # Φ) config
        hence Γk Ψk Φk k. ((((# K2 n, # K1 n  (λ(x, y). x  y)) # Γ), n
                                   Ψ  ((K1 weakly precedes K2) # Φ))
                             ↪⇗k(Γk, Suc n  Ψk  Φk))
                           (ρ   Γk, Suc n  Ψk  Φk config)
          using  WeaklyPrecedes.prems by simp
        from this obtain Γk Ψk Φk k
          where fp:(((# K2 n, # K1 n  (λ(x, y). x  y)) # Γ), n
                                   Ψ  ((K1 weakly precedes K2) # Φ))
                             ↪⇗k(Γk, Suc n  Ψk  Φk)
            and rc:ρ   Γk, Suc n  Ψk  Φk config by blast
        have (Γ, n  ((K1 weakly precedes K2) # Ψ)  Φ)
                 (((# K2 n, # K1 n  (λ(x, y). x  y)) # Γ), n
               Ψ  ((K1 weakly precedes K2) # Φ))
          by (simp add: elims_part weakly_precedes_e)
        with fp relpowp_Suc_I2 have (Γ, n  ((K1 weakly precedes K2) # Ψ)  Φ)
                                      ↪⇗Suc k(Γk, Suc n  Ψk  Φk)
          by auto
        with rc show ?thesis by blast
      qed
      ultimately show ?case using WeaklyPrecedes.prems(2) by blast
  next
    case (StrictlyPrecedes K1 K2)
      have  Γ, n  ((K1 strictly precedes K2) # Ψ)  Φ config =
         ((# K2 n, #< K1 n  (λ(x, y). x  y)) # Γ), n
           Ψ  ((K1 strictly precedes K2) # Φ) config
        using HeronConf_interp_stepwise_strictly_precedes_cases by simp
      moreover have ρ   ((# K2 n, #< K1 n  (λ(x, y). x  y)) # Γ), n
                             Ψ  ((K1 strictly precedes K2) # Φ) config
             (Γk Ψk Φk k. ((Γ, n  ((K1 strictly precedes K2) # Ψ)  Φ)
                                  ↪⇗k(Γk, Suc n  Ψk  Φk))
                 (ρ   Γk, Suc n  Ψk  Φk config))
      proof -
        assume ρ   ((# K2 n, #< K1 n  (λ(x, y). x  y)) # Γ), n
                         Ψ  ((K1 strictly precedes K2) # Φ) config
        hence Γk Ψk Φk k. ((((# K2 n, #< K1 n  (λ(x, y). x  y)) # Γ), n
                                   Ψ  ((K1 strictly precedes K2) # Φ))
                             ↪⇗k(Γk, Suc n  Ψk  Φk))
                             (ρ   Γk, Suc n  Ψk  Φk config)
          using  StrictlyPrecedes.prems by simp
        from this obtain Γk Ψk Φk k
          where fp:(((# K2 n, #< K1 n  (λ(x, y). x  y)) # Γ), n
                                   Ψ  ((K1 strictly precedes K2) # Φ))
                             ↪⇗k(Γk, Suc n  Ψk  Φk)
            and rc:ρ   Γk, Suc n  Ψk  Φk config by blast
        have (Γ, n  ((K1 strictly precedes K2) # Ψ)  Φ)
                 (((# K2 n, #< K1 n  (λ(x, y). x  y)) # Γ), n
               Ψ  ((K1 strictly precedes K2) # Φ))
          by (simp add: elims_part strictly_precedes_e)
        with fp relpowp_Suc_I2 have (Γ, n  ((K1 strictly precedes K2) # Ψ)  Φ)
                                      ↪⇗Suc k(Γk, Suc n  Ψk  Φk)
          by auto
        with rc show ?thesis by blast
      qed
      ultimately show ?case using StrictlyPrecedes.prems(2) by blast
  next
    case (Kills K1 K2)
      have branches:  Γ, n  ((K1 kills K2) # Ψ)  Φ config
          =  ((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 kills K2) # Φ) config
            ((K1  n) # (K2 ¬⇑  n) # Γ), n  Ψ  ((K1 kills K2) # Φ) config
        using HeronConf_interp_stepwise_kills_cases by simp
      moreover have br1: ρ   ((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 kills K2) # Φ) config
                 Γk Ψk Φk k. ((Γ, n  ((K1 kills K2) # Ψ)  Φ)
                                    ↪⇗k(Γk, Suc n  Ψk  Φk))
                   ρ   Γk, Suc n  Ψk  Φk config
      proof -
        assume h1: ρ   ((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 kills K2) # Φ) config
        then have Γk Ψk Φk k.
                    ((((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 kills K2) # Φ))
                    ↪⇗k(Γk, Suc n  Ψk  Φk))
                   ρ   Γk, Suc n  Ψk  Φk config
          using h1 Kills.prems by simp
        from this obtain Γk Ψk Φk k where
          fp:((((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 kills K2) # Φ))
                ↪⇗k(Γk, Suc n  Ψk  Φk))
          and rc:ρ   Γk, Suc n  Ψk  Φk config by blast
        have pc:(Γ, n  (K1 kills K2) # Ψ  Φ)
                   (((K1 ¬⇑ n) # Γ), n  Ψ  ((K1 kills K2) # Φ))
          by (simp add: elims_part kills_e1)
        hence (Γ, n  (K1 kills K2) # Ψ  Φ) ↪⇗Suc k(Γk, Suc n  Ψk  Φk)
          using fp relpowp_Suc_I2 by auto
        with rc show ?thesis by blast
      qed
      moreover have br2:
        ρ   ((K1  n) # (K2 ¬⇑  n) # Γ), n  Ψ  ((K1 kills K2) # Φ) config
           Γk Ψk Φk k. ((Γ, n  ((K1 kills K2) # Ψ)  Φ)
                                ↪⇗k(Γk, Suc n  Ψk  Φk))
                           ρ   Γk, Suc n  Ψk  Φk config
      proof -
        assume h2: ρ  ((K1  n)#(K2 ¬⇑  n)#Γ), n  Ψ  ((K1 kills K2)#Φ)config
        then have Γk Ψk Φk k. (
                    (((K1  n) # (K2 ¬⇑  n) # Γ), n  Ψ  ((K1 kills K2) # Φ))
                      ↪⇗k(Γk, Suc n  Ψk  Φk)
                    )  ρ   Γk, Suc n  Ψk  Φk config
          using h2 Kills.prems by simp
        from this obtain Γk Ψk Φk k where
            fp:(((K1  n) # (K2 ¬⇑  n) # Γ), n  Ψ  ((K1 kills K2) # Φ))
                ↪⇗k(Γk, Suc n  Ψk  Φk)
        and rc:ρ   Γk, Suc n  Ψk  Φk config by blast
        have (Γ, n  ((K1 kills K2) # Ψ)  Φ)
               (((K1  n) # (K2 ¬⇑  n) # Γ), n  Ψ  ((K1 kills K2) # Φ))
          by (simp add: elims_part kills_e2)
        hence (Γ, n  ((K1 kills K2) # Ψ)  Φ) ↪⇗Suc k(Γk, Suc n  Ψk  Φk)
          using fp relpowp_Suc_I2 by auto
        with rc show ?thesis by blast
      qed
      ultimately show ?case using Kills.prems(2) by blast
  qed
qed

lemma instant_index_increase_generalized:
  assumes n < nk
  assumes ρ   Γ, n  Ψ  Φ config
  shows   Γk Ψk Φk k. ((Γ, n  Ψ  Φ) ↪⇗k(Γk, nk  Ψk  Φk))
                          ρ   Γk, nk  Ψk  Φk config
proof -
  obtain δk where diff: nk = δk + Suc n
    using add.commute assms(1) less_iff_Suc_add by auto
  show ?thesis
    proof (subst diff, subst diff, insert assms(2), induct δk)
      case 0  thus ?case
        using instant_index_increase assms(2) by simp
    next
      case (Suc δk)
        have f0: ρ   Γ, n  Ψ  Φ config  Γk Ψk Φk k.
                  ((Γ, n  Ψ  Φ) ↪⇗k(Γk, δk + Suc n  Ψk  Φk))
                 ρ   Γk, δk + Suc n  Ψk  Φk config
          using Suc.hyps by blast
        obtain Γk Ψk Φk k
          where cont: ((Γ, n  Ψ  Φ) ↪⇗k(Γk, δk + Suc n  Ψk  Φk))
                       ρ   Γk, δk + Suc n  Ψk  Φk config
          using f0 assms(1) Suc.prems by blast
        then have fcontinue: Γk' Ψk' Φk' k'. ((Γk, δk + Suc n  Ψk  Φk)
                                       ↪⇗k'(Γk', Suc (δk + Suc n)  Ψk'  Φk'))
                                    ρ   Γk', Suc (δk + Suc n)  Ψk'  Φk' config
          using f0 cont instant_index_increase by blast
        obtain Γk' Ψk' Φk' k'
          where cont2: ((Γk, δk + Suc n  Ψk  Φk)
                        ↪⇗k'(Γk', Suc (δk + Suc n)  Ψk'  Φk'))
                       ρ   Γk', Suc (δk + Suc n)  Ψk'  Φk' config
          using Suc.prems using fcontinue cont by blast
        have trans: (Γ, n  Ψ  Φ) ↪⇗k + k'(Γk', Suc (δk + Suc n)  Ψk'  Φk')
          using operational_semantics_trans_generalized cont cont2 by blast
        moreover have suc_assoc: Suc δk + Suc n = Suc (δk + Suc n) by arith
        ultimately show ?case
          proof (subst suc_assoc)
            show Γk Ψk Φk k.
                   ((Γ, n  Ψ  Φ) ↪⇗k(Γk, Suc (δk + Suc n)  Ψk  Φk))
                   ρ   Γk, Suc δk + Suc n  Ψk  Φk config
            using cont2 local.trans by auto
          qed
    qed
qed

text‹
  Any run that belongs to a specification @{term Ψ} has a corresponding 
  configuration that develops it up to the @{term n}\textsuperscript{th} instant.
›
theorem progress:
  assumes ρ  ⟦⟦ Ψ ⟧⟧TESL
    shows k Γk Ψk Φk. (([], 0  Ψ  [])  ↪⇗k(Γk, n  Ψk  Φk))
                          ρ   Γk, n  Ψk  Φk config
proof -
  have 1:Γk Ψk Φk k. (([], 0  Ψ  []) ↪⇗k(Γk, 0  Ψk  Φk))
                       ρ   Γk, 0  Ψk  Φk config
    using assms relpowp_0_I solve_start by fastforce
  show ?thesis
  proof (cases n = 0)
    case True
      thus ?thesis using assms relpowp_0_I solve_start by fastforce
  next
    case False hence pos:n > 0 by simp
      from assms solve_start have ρ   [], 0  Ψ  [] config by blast
      from instant_index_increase_generalized[OF pos this] show ?thesis by blast
  qed
qed

section ‹Local termination›
text ‹
  Here, we prove that the computation of an instant in a run always terminates.
  Since this computation terminates when the list of constraints for the present
  instant becomes empty, we introduce a measure for this formula.
›
primrec measure_interpretation :: ::linordered_field TESL_formula  nat (μ)
where
  μ [] = (0::nat)
| μ (φ # Φ) = (case φ of
                      _ sporadic _ on _  1 + μ Φ
                    | _                  2 + μ Φ)

fun measure_interpretation_config :: ::linordered_field config  nat (μconfig)
where
  μconfig (Γ, n  Ψ  Φ) = μ Ψ

text ‹
  We then show that the elimination rules make this measure decrease.
›
lemma elimation_rules_strictly_decreasing:
  assumes (Γ1, n1  Ψ1  Φ1)  e  (Γ2, n2  Ψ2  Φ2)
    shows μ Ψ1 > μ Ψ2
using assms by (auto elim: operational_semantics_elim.cases)

lemma elimation_rules_strictly_decreasing_meas:
  assumes (Γ1, n1  Ψ1  Φ1)  e  (Γ2, n2  Ψ2  Φ2)
    shows (Ψ2, Ψ1)  measure μ
using assms by (auto elim: operational_semantics_elim.cases)

lemma elimation_rules_strictly_decreasing_meas':
  assumes 𝒮1  e  𝒮2
  shows (𝒮2, 𝒮1)  measure μconfig
proof -
  from assms obtain Γ1 n1 Ψ1 Φ1 where p1:𝒮1 = (Γ1, n1  Ψ1  Φ1)
    using measure_interpretation_config.cases by blast
  from assms obtain Γ2 n2 Ψ2 Φ2 where p2:𝒮2 = (Γ2, n2  Ψ2  Φ2)
    using measure_interpretation_config.cases by blast
  from elimation_rules_strictly_decreasing_meas assms p1 p2
    have (Ψ2, Ψ1)  measure μ by blast
  hence μ Ψ2 < μ Ψ1 by simp
  hence μconfig (Γ2, n2  Ψ2  Φ2) < μconfig (Γ1, n1  Ψ1  Φ1) by simp
  with p1 p2 show ?thesis by simp
qed

text ‹
  Therefore, the relation made up of elimination rules is well-founded and the
  computation of an instant terminates.
›
theorem instant_computation_termination:
  wfP (λ(𝒮1::'a::linordered_field config) 𝒮2. (𝒮1  e  𝒮2))
proof (simp add: wfP_def)
  show wf {((𝒮1::'a::linordered_field config), 𝒮2). 𝒮1 e 𝒮2}
  proof (rule wf_subset)
    have measure μconfig = {(𝒮2, (𝒮1::'a::linordered_field config)).
                               μconfig 𝒮2 < μconfig 𝒮1}
      by (simp add: inv_image_def less_eq measure_def)
    thus {((𝒮1::'a::linordered_field config), 𝒮2). 𝒮1 e 𝒮2}  (measure μconfig)
      using elimation_rules_strictly_decreasing_meas'
            operational_semantics_elim_inv_def by blast
  next
    show wf (measure measure_interpretation_config) by simp
  qed
qed


end