Theory SmallStepCon

theory SmallStepCon
imports SmallStep TerminationCon
(*
    Author:      David Sanan
    Maintainer:  David Sanan, sanan at ntu edu sg
    License:     LGPL
*)

(*  Title:      SmallStepCon.thy
    Author:     David Sanan, NTU

Copyright (C) 2015-2016 David Sanan 
Some rights reserved, NTU
This library is free software; you can redistribute it and/or modify
it under the terms of the GNU Lesser General Public License as
published by the Free Software Foundation; either version 2.1 of the
License, or (at your option) any later version.

This library is distributed in the hope that it will be useful, but
WITHOUT ANY WARRANTY; without even the implied warranty of
MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU
Lesser General Public License for more details.

You should have received a copy of the GNU Lesser General Public
License along with this library; if not, write to the Free Software
Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307
USA
*)

section {* Small-Step Semantics and Infinite Computations*} 

theory SmallStepCon imports "EmbSimpl/SmallStep" "SemanticCon"  
                            "TerminationCon"
                            (* "Sep_Algebra.Sep_Heap_Instance" 
                            "Actions.ActionsSemantics" *)
begin

text {* The redex of a statement is the substatement, which is actually altered
by the next step in the small-step semantics.
*}
primrec redex:: "('s,'p,'f,'e)com ⇒ ('s,'p,'f,'e)com"
where
"redex Skip = Skip" |
"redex (Basic f e) = (Basic f e)" |
"redex (Spec r e) = (Spec r e)" |
"redex (Seq c1 c2) = redex c1" |
"redex (Cond b c1 c2) = (Cond b c1 c2)" |
"redex (While b c) = (While b c)" |
"redex (Call p) = (Call p)" |
"redex (DynCom d) = (DynCom d)" |
"redex (Guard f b c) = (Guard f b c)" |
"redex (Throw) = Throw" |
"redex (Catch c1 c2) = redex c1" |
"redex (Await b c e) = (Await b c e)"



subsection {*Small-Step Computation: @{text "Γ⊢c(c, s) → (c', s')"}*}
type_synonym ('s,'p,'f,'e) config = "('s,'p,'f,'e)com  × ('s,'f) xstate"

inductive 
      "step_e"::"[('s,'p,'f,'e) body,('s,'p,'f,'e) config,('s,'p,'f,'e) config] ⇒ bool"
                                ("_⊢c (_ →e/ _)" [81,81,81] 100)  
  for Γ::"('s,'p,'f,'e) body"
where
Env: "Γ⊢c (Ps, Normal s) →e (Ps, t)"
|Env_n: "(∀t'. t≠Normal t') ⟹ Γ⊢c (Ps, t) →e (Ps, t)"

lemma etranE: "Γ⊢c c →e c' ⟹ (⋀P s t. c = (P, s) ⟹ c' = (P, t) ⟹ Q) ⟹ Q"
  by (induct c, induct c', erule  step_e.cases, blast)

inductive_cases stepe_Normal_elim_cases [cases set]:
 "Γ⊢c(Ps,Normal s) →e (Ps,t)"

inductive_cases stepe_elim_cases [cases set]:
 "Γ⊢c(Ps,s) →e (Ps,t)"

inductive_cases stepe_not_norm_elim_cases [cases set]:
 "Γ⊢c(Ps,s) →e (Ps,Abrupt t)"
 "Γ⊢c(Ps,s) →e (Ps,Stuck)"
 "Γ⊢c(Ps,s) →e (Ps,Fault t)"
 "Γ⊢c(Ps,s) →e (Ps,Normal t)"

lemma env_c_c'_false:
   assumes step_m: "Γ⊢c (c, s) →e (c', s')" 
   shows "~(c=c')  ⟹ P"
using step_m etranE by blast

lemma eenv_normal_s'_normal_s:
   assumes step_m: "Γ⊢c (c, s) →e (c', Normal s')" 
   shows "(⋀s1. s≠Normal s1)  ⟹ P"
using step_m 
by (cases, auto)

lemma env_normal_s'_normal_s:
   assumes step_m: "Γ⊢c (c, s) →e (c', Normal s') " 
   shows "∃s1. s= Normal s1"
using step_m 
by (cases, auto)

lemma env_c_c':
   assumes step_m: "Γ⊢c (c, s) →e (c', s')" 
   shows "(c=c')"
using env_c_c'_false step_m by fastforce 

lemma env_normal_s:
   assumes step_m: "Γ⊢c (c, s) →e (c', s') ∧ s≠s'" 
   shows "∃sa. s = Normal sa"
using prod.inject step_e.cases step_m by fastforce

lemma env_not_normal_s:
   assumes a1:"Γ⊢c (c, s) →e (c', s')" and  a2:"(∀t. s≠Normal t)" 
   shows "s=s'"
using a1 a2
by (cases rule:step_e.cases,auto) 

lemma env_not_normal_s_not_norma_t:
   assumes a1:"Γ⊢c (c, s) →e (c', s')" and  a2:"(∀t. s≠Normal t)" 
   shows "(∀t. s'≠Normal t)"
using a1 a2 env_not_normal_s
by blast

lemma stepe_not_Fault_f_end:
  assumes step_e: "Γ⊢c (c1, s) →e (c1', s')"
  shows "s'∉ Fault ` f ⟹ s ∉ Fault ` f"
proof (cases s) 
  case (Fault f')
    assume s'_f:"s' ∉ Fault ` f" and
           "s = Fault f'" 
    then have "s=s'" using step_e 
    using env_normal_s xstate.distinct(3) by blast  
  thus ?thesis using s'_f Fault by blast
qed (auto)

inductive       
      "stepc"::"[('s,'p,'f,'e) body,('s,'p,'f,'e) config,('s,'p,'f,'e) config] ⇒ bool"
                                ("_⊢c (_ →/ _)" [81,81,81] 100)  
  for Γ::"('s,'p,'f,'e) body"
where

  Basicc: "Γ⊢c(Basic f e,Normal s) → (Skip,Normal (f s))"

| Specc: "(s,t) ∈ r ⟹ Γ⊢c(Spec r e,Normal s) → (Skip,Normal t)"
| SpecStuckc: "∀t. (s,t) ∉ r ⟹ Γ⊢c(Spec r e,Normal s) → (Skip,Stuck)"

| Guardc: "s∈g ⟹ Γ⊢c(Guard f g c,Normal s) → (c,Normal s)"
  
| GuardFaultc: "s∉g ⟹ Γ⊢c(Guard f g c,Normal s) → (Skip,Fault f)"


| Seqc: "Γ⊢c(c1,s) → (c1',s')
        ⟹ 
        Γ⊢c(Seq c1 c2,s) → (Seq c1' c2, s')"
| SeqSkipc: "Γ⊢c(Seq Skip c2,s) → (c2, s)"
| SeqThrowc: "Γ⊢c(Seq Throw c2,Normal s) → (Throw, Normal s)"

| CondTruec:  "s∈b ⟹ Γ⊢c(Cond b c1 c2,Normal s) → (c1,Normal s)"
| CondFalsec: "s∉b ⟹ Γ⊢c(Cond b c1 c2,Normal s) → (c2,Normal s)"

| WhileTruec: "⟦s∈b⟧ 
              ⟹ 
              Γ⊢c(While b c,Normal s) → (Seq c (While b c),Normal s)"

| WhileFalsec: "⟦s∉b⟧ 
               ⟹ 
               Γ⊢c(While b c,Normal s) → (Skip,Normal s)"


| Awaitc:  "⟦s∈b; Γ1=Γ¬a ; Γ1⊢⟨ca1,Normal s⟩ ⇒ t; 
             ¬(∃t'. t = Abrupt t')⟧ ⟹ 
            Γ⊢c(Await b ca1 e,Normal s) → (Skip,t)"

| AwaitAbruptc: "⟦s∈b; Γ1=Γ¬a ; Γ1⊢⟨ca1,Normal s⟩ ⇒ t; 
                  t = Abrupt t'⟧ ⟹ 
                 Γ⊢c(Await b ca1 e,Normal s) → (Throw,Normal t')"

| Callc: "⟦Γ p = Some bdy ; bdy≠Call p⟧ ⟹
         Γ⊢c(Call p,Normal s) → (bdy,Normal s)"

| CallUndefinedc: "Γ p=None ⟹
         Γ⊢c(Call p,Normal s) → (Skip,Stuck)"

| DynComc: "Γ⊢c(DynCom c,Normal s) → (c s,Normal s)"

| Catchc: "⟦Γ⊢c(c1,s) → (c1',s')⟧
          ⟹
          Γ⊢c(Catch c1 c2,s) → (Catch c1' c2,s')"

| CatchThrowc: "Γ⊢c(Catch Throw c2,Normal s) → (c2,Normal s)"
| CatchSkipc: "Γ⊢c(Catch Skip c2,s) → (Skip,s)"

| FaultPropc:  "⟦c≠Skip; redex c = c⟧ ⟹ Γ⊢c(c,Fault f) → (Skip,Fault f)"
| StuckPropc:  "⟦c≠Skip; redex c = c⟧ ⟹ Γ⊢c(c,Stuck) → (Skip,Stuck)"
| AbruptPropc: "⟦c≠Skip; redex c = c⟧ ⟹ Γ⊢c(c,Abrupt f) → (Skip,Abrupt f)"
                                                              
lemmas stepc_induct = stepc.induct [of _ "(c,s)" "(c',s')", split_format (complete), case_names
Basicc Specc SpecStuckc Guardc GuardFaultc Seqc SeqSkipc SeqThrowc CondTruec CondFalsec
WhileTruec WhileFalsec Awaitc AwaitAbruptc Callc CallUndefinedc DynComc Catchc CatchThrowc CatchSkipc
FaultPropc StuckPropc AbruptPropc, induct set]


inductive_cases stepc_elim_cases [cases set]:
 "Γ⊢c(Skip,s) → u"
 "Γ⊢c(Guard f g c,s) → u"
 "Γ⊢c(Basic f e,s) → u"
 "Γ⊢c(Spec r e,s) → u"
 "Γ⊢c(Seq c1 c2,s) → u"
 "Γ⊢c(Cond b c1 c2,s) → u"
 "Γ⊢c(While b c,s) → u"
 "Γ⊢c(Await b c2 e,s) → u"
 "Γ⊢c(Call p,s) → u"
 "Γ⊢c(DynCom c,s) → u"
 "Γ⊢c(Throw,s) → u"
 "Γ⊢c(Catch c1 c2,s) → u"

inductive_cases stepc_not_normal_elim_cases:
 "Γ⊢c(Call p,Abrupt s) → (p',s')"
 "Γ⊢c(Call p, Fault f) → (p',s')"
 "Γ⊢c(Call p, Stuck) → (p',s')"
 

lemma Guardc_not_c:"Guard f g c ≠ c"
proof (induct c)
qed auto

lemma Catch_not_c1:"Catch c1 c2 ≠ c1"
proof (induct c1)
qed auto

lemma Catch_not_c:"Catch c1 c2 ≠ c2"
proof (induct c2)
qed auto

lemma seq_not_eq1: "Seq c1 c2≠c1"
  by (induct c1) auto

lemma seq_not_eq2: "Seq c1 c2≠c2"
  by (induct c2) auto

lemma if_not_eq1: "Cond b c1 c2 ≠c1"
  by (induct c1) auto

lemma if_not_eq2: "Cond b c1 c2≠c2"
  by (induct c2) auto


lemmas seq_and_if_not_eq [simp] = seq_not_eq1 seq_not_eq2 
seq_not_eq1 [THEN not_sym] seq_not_eq2 [THEN not_sym] 
if_not_eq1 if_not_eq2 if_not_eq1 [THEN not_sym] if_not_eq2 [THEN not_sym]
Catch_not_c1 Catch_not_c Catch_not_c1 [THEN not_sym] Catch_not_c[THEN not_sym] 
Guardc_not_c Guardc_not_c[THEN not_sym]

inductive_cases stepc_elim_cases_Seq_Seq:
"Γ⊢c(Seq c1 c2,s) → (Seq c1' c2,s')" 

inductive_cases stepc_elim_cases_Seq_Seq1:
"Γ⊢c(Seq c1 c2,Fault f) → (q,s')" 
thm stepc_elim_cases_Seq_Seq1

inductive_cases stepc_elim_cases_Catch_Catch:
"Γ⊢c(Catch c1 c2,s) → (Catch c1' c2,s')" 

inductive_cases stepc_elim_cases_Catch_Catch1:
"Γ⊢c(Seq c1 c2,Fault f) → (q,s')" 

inductive_cases stepc_elim_cases_Seq_skip:
"Γ⊢c(Seq Skip c2,s) → u" 
"Γ⊢c(Seq (Guard f g c1) c2,s) → u"


inductive_cases stepc_elim_cases_Catch_skip:
"Γ⊢c(Catch Skip c2,s) → u"

inductive_cases stepc_elim_cases_Await_skip:
"Γ⊢c (Await b c e, Normal s) → (Skip,t)"

inductive_cases stepc_elim_cases_Await_throw:
"Γ⊢c (Await b c e, Normal s) → (Throw,t)"

inductive_cases stepc_elim_cases_Catch_throw:
"Γ⊢c(Catch c1 c2,s) → (Throw, Normal s1)" 

inductive_cases stepc_elim_cases_Catch_skip_c2:
"Γ⊢c(Catch c1 c2,s) → (c2,s)"

inductive_cases stepc_Normal_elim_cases [cases set]:
 "Γ⊢c(Skip,Normal s) → u"
 "Γ⊢c(Guard f g c,Normal s) → u"
 "Γ⊢c(Basic f e,Normal s) → u"
 "Γ⊢c(Spec r e,Normal s) → u"
 "Γ⊢c(Seq c1 c2,Normal s) → u"
 "Γ⊢c(Cond b c1 c2,Normal s) → u"
 "Γ⊢c(While b c,Normal s) → u"
 "Γ⊢c(Await b c e,Normal s) → u"
 "Γ⊢c(Call p,Normal s) → u"
 "Γ⊢c(DynCom c,Normal s) → u"
 "Γ⊢c(Throw,Normal s) → u"
 "Γ⊢c(Catch c1 c2,Normal s) → u"


text ‹ The final configuration is either of the form @{text "(Skip,_)"} for normal
termination, or @{term "(Throw,Normal s)"} in case the program was started in 
a @{term "Normal"} state and terminated abruptly. The @{const "Abrupt"} state is not used to
model abrupt termination, in contrast to the big-step semantics. Only if the
program starts in an @{const "Abrupt"} states it ends in the same @{term "Abrupt"}
state.›

definition final:: "('s,'p,'f,'e) config ⇒ bool" where
"final cfg ≡ (fst cfg=Skip ∨ ((fst cfg=Throw) ∧ (∃s. snd cfg=Normal s)))"

definition final_valid::"('s,'p,'f,'e) config ⇒ bool" where
"final_valid cfg = ((fst cfg=Skip ∨ fst cfg=Throw) ∧ (∃s. snd cfg=Normal s))"

abbreviation 
 "stepc_rtrancl" :: "[('s,'p,'f,'e) body,('s,'p,'f,'e) config,('s,'p,'f,'e) config] ⇒ bool"
                                ("_⊢c (_ →*/ _)" [81,81,81] 100)
 where                                
  "Γ⊢c cf0 →* cf1 ≡ ((CONST stepc Γ))** cf0 cf1" 
  (* "Γ⊢c cf0 →* cf1 ≡ (CONST ((stepc Γ) ∪ (step_e Γ)))** cf0 cf1" *)

abbreviation 
 "stepc_trancl" :: "[('s,'p,'f,'e) body,('s,'p,'f,'e) config,('s,'p,'f,'e) config] ⇒ bool"
                                ("_⊢c (_ →+/ _)" [81,81,81] 100)
 where
  "Γ⊢c cf0 →+ cf1 ≡ (CONST stepc Γ)++ cf0 cf1"

lemma 
   assumes 
           step_a: "Γ⊢c(Await b c e, Normal s) → (t,u)"
   shows step_await_step_c:"(Γ¬a)⊢(c, Normal s) →* (sequential t,u)" 
using step_a
proof cases
  fix t1
  assume
      "(t, u) = (Skip, t1)" "s ∈ b" "(Γ¬a)⊢ ⟨c,Normal s⟩ ⇒ t1" "∀t'. t1 ≠ Abrupt t'"
  thus ?thesis 
   by (cases u) 
   (auto intro: exec_impl_steps_Fault exec_impl_steps_Normal exec_impl_steps_Stuck)
next
  fix t1
  assume "(t, u) = (Throw, Normal t1)" "s ∈ b" "(Γ¬a)⊢ ⟨c,Normal s⟩ ⇒ Abrupt t1"
  thus ?thesis by (simp add: exec_impl_steps_Normal_Abrupt)
qed

lemma 
   assumes (* exec: "Γ⊢⟨c,Normal s⟩ ⇒ t" and
           b: "s ∈ b" and *)
           step_a: "Γ⊢c(Await b c e, Normal s) → u"
   shows step_await_final1:"final u"
using step_a 
proof cases
  case  (1 t) thus "final u"  by (simp add: final_def)
next
  case (2 t)
  thus "final u" by (simp add: exec_impl_steps_Normal_Abrupt final_def)
qed

lemma step_Abrupt_end: 
  assumes step: "Γ⊢c (c1, s) → (c1', s')"
  shows "s'=Abrupt x ⟹ s=Abrupt x"
using step
by induct auto


lemma step_Stuck_end: 
  assumes step: "Γ⊢c (c1, s) → (c1', s')"
  shows "s'=Stuck ⟹ 
          s=Stuck ∨ 
          (∃r x e. redex c1 = Spec r e ∧ s=Normal x ∧ (∀t. (x,t)∉r)) ∨ 
          (∃p x. redex c1=  Call p ∧ s=Normal x ∧ Γ p = None) ∨
          (∃b c x e.  redex c1 = Await b c e ∧ s=Normal x ∧ x ∈ b ∧ (Γ¬a)⊢⟨c,s⟩⇒s')"
using step
by induct auto

lemma step_Fault_end: 
  assumes step: "Γ⊢c (c1, s) → (c1', s')"
  shows "s'=Fault f ⟹ 
          s=Fault f ∨ 
          (∃g c x.  redex c1 = Guard f g c ∧ s=Normal x ∧ x ∉ g) ∨
          (∃b c1 x e.  redex c1 = Await b c1 e ∧ s=Normal x ∧ x ∈ b ∧ (Γ¬a)⊢⟨c1,s⟩⇒s')"
using step 
by induct auto

lemma step_not_Fault_f_end:
  assumes step: "Γ⊢c (c1, s) → (c1', s')"
  shows "s'∉ Fault ` f ⟹ s ∉ Fault ` f"
using step 
by induct auto

  

inductive 
      "step_ce"::"[('s,'p,'f,'e) body,('s,'p,'f,'e) config,('s,'p,'f,'e) config] ⇒ bool"
                                ("_⊢c (_ →ce/ _)" [81,81,81] 100)  
  for Γ::"('s,'p,'f,'e) body"
where
c_step: "Γ⊢c cf0 → cf1 ⟹ Γ⊢c cf0 →ce cf1 "
|e_step: "Γ⊢c cf0 →e cf1 ⟹ Γ⊢c cf0 →ce cf1 "

lemmas step_ce_induct = step_ce.induct [of _ "(c,s)" "(c',s')", split_format (complete), case_names
c_step e_step, induct set]


inductive_cases step_ce_elim_cases [cases set]:
 "Γ⊢c cf0 →ce cf1"


lemma step_c_normal_normal: assumes a1: "Γ⊢c cf0 → cf1"
      shows "⋀ c1 s s'. ⟦cf0 = (c1,Normal s);cf1=(c1,s');(∀sa. ¬(s'=Normal sa))⟧
          ⟹ P"
using a1 
by (induct rule: stepc.induct, induct, auto)

lemma normal_not_normal_eq_p: 
  assumes a1: "Γ⊢c cf0 →ce cf1"
  shows "⋀ c1 s s'. ⟦cf0 = (c1,Normal s);cf1=(c1,s');(∀sa. ¬(s'=Normal sa))⟧
          ⟹ Γ⊢c cf0 →e cf1 ∧ ¬( Γ⊢c cf0 → cf1)"
by (meson step_c_normal_normal step_e.intros)

lemma call_not_normal_skip_always:
  assumes a0:"Γ⊢c(Call p,s) → (p1,s1)" and
          a1:"∀sn. s ≠ Normal sn" and
          a2:"p1≠Skip"
  shows "P" 
proof(cases s)
  case Normal thus ?thesis using a1 by fastforce
next
  case Stuck 
  then have a0:"Γ⊢c(Call p,Stuck) → (p1,s1)" using a0 by auto
  show ?thesis using  a1 a2 stepc_not_normal_elim_cases(3)[OF a0] by fastforce
next
  case (Fault f) 
  then have a0:"Γ⊢c(Call p,Fault f) → (p1,s1)" using a0 by auto
  show ?thesis using  a1 a2 stepc_not_normal_elim_cases(2)[OF a0] by fastforce
next
  case (Abrupt a)
  then have a0:"Γ⊢c(Call p,Abrupt a) → (p1,s1)" using a0 by auto
  show ?thesis using  a1 a2 stepc_not_normal_elim_cases(1)[OF a0] by fastforce
qed

lemma call_f_step_not_s_eq_t_false:
  assumes 
     a0:"Γ⊢c(P,s) → (Q,t)" and
     a1:"(redex P = Call fn ∧ Γ fn = Some bdy ∧ s=Normal s' ∧ ~(s=t)) ∨
         (redex P = Call fn ∧ Γ fn = Some bdy ∧ s=Normal s' ∧ s=t ∧ P=Q ∧ Γ fn ≠ Some (Call fn))"
  shows "False"
using a0 a1
proof (induct rule:stepc_induct)
qed(fastforce+,auto)

lemma call_f_step_ce_not_s_eq_t_env_step:
  assumes 
     a0:"Γ⊢c(P,s) →ce (Q,t)" and
     a1:"(redex P = Call fn ∧ Γ fn = Some bdy ∧ s=Normal s' ∧ ~(s=t)) ∨
         (redex P = Call fn ∧ Γ fn = Some bdy ∧ s=Normal s' ∧ s=t ∧ P=Q ∧ Γ fn ≠ Some (Call fn)) "
  shows "Γ⊢c(P,s) →e (Q,t)"
proof-
  have "Γ⊢c(P,s) →e (Q,t) ∨ Γ⊢c(P,s) → (Q,t)"
  using a0 step_ce_elim_cases by fastforce
  thus ?thesis using call_f_step_not_s_eq_t_false a1 by fastforce
qed


  
abbreviation 
 "stepce_rtrancl" :: "[('s,'p,'f,'e) body,('s,'p,'f,'e) config,('s,'p,'f,'e) config] ⇒ bool"
                                ("_⊢c (_ →ce*/ _)" [81,81,81] 100)
 where                                
  "Γ⊢c cf0 →ce* cf1 ≡ ((CONST step_ce Γ))** cf0 cf1" 
  (* "Γ⊢c cf0 →* cf1 ≡ (CONST ((stepc Γ) ∪ (step_e Γ)))** cf0 cf1" *)

abbreviation 
 "stepce_trancl" :: "[('s,'p,'f,'e) body,('s,'p,'f,'e) config,('s,'p,'f,'e) config] ⇒ bool"
                                ("_⊢c (_ →ce+/ _)" [81,81,81] 100)
 where
  "Γ⊢c cf0 →ce+ cf1 ≡ (CONST step_ce Γ)++ cf0 cf1"

subsection {*Parallel Computation: @{text "Γ⊢(c, s) →p  (c', s')"}*}
type_synonym ('s,'p,'f,'e) par_Simpl = "('s,'p,'f,'e)com  list" 
type_synonym ('s,'p,'f,'e) par_config = "('s,'p,'f,'e) par_Simpl × ('s,'f) xstate"

definition final_c:: "('s,'p,'f,'e) par_config ⇒ bool" where
"final_c cfg = (∀i. i<length (fst cfg) ⟶ final ((fst cfg)!i, snd cfg))"

inductive 
      "step_pe"::"[('s,'p,'f,'e) body,('s,'p,'f,'e) par_config,('s,'p,'f,'e) par_config] ⇒ bool"
                                ("_⊢p (_ →e/ _)" [81,81,81] 100)  
  for Γ::"('s,'p,'f,'e) body"
where
ParEnv: "Γ⊢p (Ps, Normal s) →e (Ps, Normal t)"




lemma ptranE: "Γ⊢p c →e c' ⟹ (⋀P s t. c = (P, s) ⟹ c' = (P, t) ⟹ Q) ⟹ Q"
  by (induct c, induct c', erule  step_pe.cases, blast)

inductive_cases step_pe_Normal_elim_cases [cases set]:
 "Γ⊢p(PS,Normal s) →e (Ps,t)"

inductive_cases step_pe_elim_cases [cases set]:
 "Γ⊢p(PS,s) →e (Ps,t)"

inductive_cases step_pe_not_norm_elim_cases [cases set]:
 "Γ⊢p(Ps,s) →e (Ps,Abrupt t)"
 "Γ⊢p(Ps,s) →e (Ps,Stuck)"
 "Γ⊢p(Ps,s) →e (Ps,Fault t)"

lemma env_pe_c_c'_false:
   assumes step_m: "Γ⊢p (c, s) →e (c', s')" 
   shows "~(c=c')  ⟹ P"
using step_m ptranE by blast

lemma env_pe_c_c':
   assumes step_m: "Γ⊢p (c, s) →e (c', s')" 
   shows "(c=c')"
using env_pe_c_c'_false step_m by fastforce 

lemma env_pe_normal_s:
   assumes step_m: "Γ⊢p (c, s) →e (c', s') ∧ s≠s'" 
   shows "∃sa. s = Normal sa"
using prod.inject step_pe.cases step_m by fastforce

lemma env_pe_not_normal_s:
   assumes a1:"Γ⊢p (c, s) →e (c', s')" and  a2:"(∀t. s≠Normal t)" 
   shows "s=s'"
using a1 a2
by (cases rule:step_pe.cases,auto) 

lemma env_pe_not_normal_s_not_norma_t:
   assumes a1:"Γ⊢p (c, s) →e (c', s')" and  a2:"(∀t. s≠Normal t)" 
   shows "(∀t. s'≠Normal t)"
using a1 a2 env_pe_not_normal_s
by blast


inductive       
"step_p"::"[('s,'p,'f,'e) body, ('s,'p,'f,'e) par_config, 
            ('s,'p,'f,'e) par_config] ⇒ bool"
("_⊢p (_ →/ _)" [81,81,81] 100)  
where
 ParComp: "⟦i<length Ps;  Γ⊢c(Ps!i,s) → (r,s')⟧ ⟹  
           Γ⊢p(Ps, s) → (Ps[i:=r], s')"

lemmas steppe_induct = step_p.induct [of _ "(c,s)" "(c',s')", split_format (complete), case_names
ParComp, induct set]

inductive_cases step_p_elim_cases [cases set]:
 "Γ⊢p(Ps, s) → u"

inductive_cases step_p_pair_elim_cases [cases set]:
 "Γ⊢p(Ps, s) → (Qs, t)"

inductive_cases step_p_Normal_elim_cases [cases set]:
 "Γ⊢p(Ps, Normal s) → u"


lemma par_ctranE: "Γ ⊢p c → c' ⟹
  (⋀i Ps s r t. c = (Ps, s) ⟹ c' = (Ps[i := r], t) ⟹ i < length Ps ⟹
     Γ ⊢c (Ps!i, s) → (r, t) ⟹ P) ⟹ P"
by (induct c, induct c', erule step_p.cases, blast)

subsection ‹Computations›
subsubsection ‹Sequential computations›

type_synonym ('s,'p,'f,'e) confs = 
  "('s,'p,'f,'e) body ×(('s,'p,'f,'e) config) list"

inductive_set cptn :: "(('s,'p,'f,'e) confs) set"
where
  CptnOne: " (Γ, [(P,s)]) ∈ cptn"
| CptnEnv: "⟦Γ⊢c(P,s) →e (P,t); (Γ,(P, t)#xs) ∈ cptn ⟧ ⟹
             (Γ,(P,s)#(P,t)#xs) ∈ cptn"
| CptnComp: "⟦Γ⊢c(P,s) → (Q,t); (Γ,(Q, t)#xs) ∈ cptn ⟧ ⟹ 
              (Γ,(P,s)#(Q,t)#xs) ∈ cptn"

inductive_cases cptn_elim_cases [cases set]:
"(Γ, [(P,s)]) ∈ cptn"
"(Γ,(P,s)#(Q,t)#xs) ∈ cptn"
"(Γ,(P,s)#(P,t)#xs) ∈ cptn"

inductive_cases cptn_elim_cases_pair [cases set]:
"(Γ, [x]) ∈ cptn"
"(Γ, x#x1#xs) ∈ cptn"

lemma cptn_dest:"(Γ,(P,s)#(Q,t)#xs) ∈ cptn ⟹ (Γ,(Q,t)#xs)∈ cptn"
by (auto dest: cptn_elim_cases)

lemma cptn_dest_pair:"(Γ,x#x1#xs) ∈ cptn ⟹ (Γ,x1#xs)∈ cptn"
proof -
  assume "(Γ,x#x1#xs) ∈ cptn"
  thus ?thesis using cptn_dest prod.collapse by metis
qed

lemma cptn_dest1:"(Γ,(P,s)#(Q,t)#xs) ∈ cptn ⟹ (Γ,(P,s)#[(Q,t)])∈ cptn"
proof -
  assume a1: "(Γ, (P, s) # (Q, t) # xs) ∈ cptn"
  have "(Γ, [(Q, t)]) ∈ cptn"
    by (meson cptn.CptnOne)
  thus ?thesis       
  proof (cases s)
    case (Normal s') 
     then have f1: "(Γ, (P, Normal s') # (Q, t) # xs) ∈ cptn"
       using Normal a1 by blast
     have "(Γ, [(P, t)]) ∈ cptn ⟶ (Γ, [(P, Normal s'), (P, t)]) ∈ cptn"
       by (simp add: Env cptn.CptnEnv)       
     thus ?thesis
      using f1 by (metis (no_types) Normal `(Γ, [(Q, t)]) ∈ cptn` cptn.CptnComp cptn_elim_cases(2))
  next
    case (Abrupt x) thus ?thesis  
     using `(Γ, [(Q, t)]) ∈ cptn` a1 cptn.CptnComp cptn_elim_cases(2) CptnEnv by metis
  next
    case (Stuck) thus ?thesis 
     using `(Γ, [(Q, t)]) ∈ cptn` a1 cptn.CptnComp cptn_elim_cases(2) CptnEnv by metis
  next 
    case (Fault f) thus ?thesis
     using `(Γ, [(Q, t)]) ∈ cptn` a1 cptn.CptnComp cptn_elim_cases(2) CptnEnv by metis
  qed
qed

lemma cptn_dest1_pair:"(Γ,x#x1#xs) ∈ cptn ⟹ (Γ,x#[x1])∈ cptn"
proof -
  assume "(Γ,x#x1#xs) ∈ cptn"
  thus ?thesis using cptn_dest1 prod.collapse by metis
qed

lemma cptn_append_is_cptn [rule_format]: 
 "∀b a. (Γ,b#c1)∈cptn ⟶  (Γ,a#c2)∈cptn ⟶ (b#c1)!length c1=a ⟶ (Γ,b#c1@c2)∈cptn"
apply(induct c1)
 apply simp
apply clarify
apply(erule cptn.cases,simp_all)
apply (simp add: cptn.CptnEnv)
by (simp add: cptn.CptnComp)

lemma cptn_dest_2:
  "(Γ,a#xs@ys) ∈ cptn  ⟹ (Γ,a#xs)∈ cptn"
proof (induct "xs" arbitrary: a)
  case Nil thus ?case using cptn.simps by fastforce 
next
  case (Cons x xs') 
  then have "(Γ,a#[x])∈cptn" by (simp add: cptn_dest1_pair)
  also have "(Γ, x # xs') ∈ cptn"
    using Cons.hyps Cons.prems cptn_dest_pair by fastforce    
  ultimately show ?case using cptn_append_is_cptn [of Γ a "[x]" x xs']
    by force    
qed   


lemma last_not_F:
assumes 
  a0:"(Γ,xs)∈cptn"  
shows "snd (last xs) ∉ Fault ` F ⟹ ∀i < length xs. snd (xs!i) ∉ Fault ` F"
using a0
proof(induct) print_cases
  case (CptnOne Γ p s) thus ?case by auto
next
  case (CptnEnv  Γ P s t xs) 
  thus ?case using stepe_not_Fault_f_end
  proof -
  { fix nn :: nat
    have "snd (last ((P, t) # xs)) ∉ Fault ` F"
      using CptnEnv.prems by force
    then have "¬ nn < length ((P, s) # (P, t) # xs) ∨ snd (((P, s) # (P, t) # xs) ! nn) ∉ Fault ` F"
      by (metis (no_types) CptnEnv.hyps(1) CptnEnv.hyps(3) length_Cons less_Suc_eq_0_disj nth_Cons_0 nth_Cons_Suc snd_conv stepe_not_Fault_f_end) 
  }
  then have "∀n. ¬ n < length ((P, s) # (P, t) # xs) ∨ snd (((P, s) # (P, t) # xs) ! n) ∉ Fault ` F"
    by meson
  then show ?thesis
    by metis
  qed    
next
  case (CptnComp Γ P s Q t xs) 
  have "snd (last ((Q, t) # xs)) ∉ Fault ` F"
    using CptnComp.prems by force
  then have all:"∀i<length ((Q, t) # xs). snd (((Q, t) # xs) ! i) ∉ Fault ` F"
    using CptnComp.hyps by force
  then have "t ∉ Fault ` F"
    by force
  then have "s ∉ Fault ` F" using step_not_Fault_f_end
    using CptnComp.hyps(1) by blast
  then have zero:"snd (P,s) ∉ Fault ` F" by auto
  show ?case 
  proof -
  { fix nn :: nat
    have "¬ nn < length ((P, s) # (Q, t) # xs) ∨ snd (((P, s) # (Q, t) # xs) ! nn) ∉ Fault ` F"
      by (metis (no_types) ‹∀i<length ((Q, t) # xs). snd (((Q, t) # xs) ! i) ∉ Fault ` F› ‹snd (P, s) ∉ Fault ` F› diff_Suc_1 length_Cons less_Suc_eq_0_disj nth_Cons') 
  }
  then show ?thesis
    by meson
  qed 
qed

definition cp :: "('s,'p,'f,'e) body ⇒ ('s,'p,'f,'e) com ⇒ 
                  ('s,'f) xstate ⇒ (('s,'p,'f,'e) confs) set" where
  "cp Γ P s ≡ {(Γ1,l). l!0=(P,s) ∧ (Γ,l) ∈ cptn ∧ Γ1=Γ}"  

 

lemma cp_sub: 
  assumes a0: "(Γ,(x#l0)@l1) ∈ cp Γ P s"
  shows "(Γ,(x#l0)) ∈ cp Γ P s"
proof -
  have "(x#l0)!0 = (P,s)" using a0 unfolding cp_def by auto
  also have "(Γ,(x#l0))∈cptn" using a0 unfolding cp_def
  using cptn_dest_2 by fastforce
  ultimately show ?thesis using a0 unfolding cp_def by blast
qed

subsubsection ‹Parallel computations›

type_synonym ('s,'p,'f,'e) par_confs = "('s,'p,'f,'e) body ×(('s,'p,'f,'e) par_config) list"

inductive_set par_cptn :: "('s,'p,'f,'e) par_confs set"
where
  ParCptnOne: "(Γ, [(P,s)]) ∈ par_cptn"
| ParCptnEnv: "⟦Γ⊢p(P,s) →e (P,t);(Γ,(P, t)#xs) ∈ par_cptn ⟧ ⟹(Γ,(P,s)#(P,t)#xs) ∈ par_cptn"
| ParCptnComp: "⟦ Γ ⊢p(P,s) → (Q,t); (Γ,(Q,t)#xs) ∈ par_cptn ⟧ ⟹ (Γ,(P,s)#(Q,t)#xs) ∈ par_cptn"

inductive_cases par_cptn_elim_cases [cases set]:
"(Γ, [(P,s)]) ∈ par_cptn"
"(Γ,(P,s)#(Q,t)#xs) ∈ par_cptn"

lemma pe_ce: 
  assumes a1:"Γ⊢p(P,s) →e (P,t)"
  shows "∀i<length P. Γ⊢c(P!i,s) →e (P!i,t)"
proof -
  {fix i
   assume "i< length P"
   have "Γ⊢c(P!i,s) →e (P!i,t)" using a1
  by (metis Env Env_n env_pe_not_normal_s)
  }
  thus "∀i<length P. Γ⊢c(P!i,s) →e (P!i,t)" by blast
qed

type_synonym ('s,'p,'f,'e) par_com = "('s,'p,'f,'e) com list"

definition par_cp :: "('s,'p,'f,'e) body ⇒ ('s,'p,'f,'e) com list ⇒ ('s,'f) xstate ⇒ (('s,'p,'f,'e) par_confs) set" 
where
  "par_cp Γ P s ≡ {(Γ1,l). l!0=(P,s) ∧ (Γ,l) ∈ par_cptn ∧ Γ1=Γ}"  

(* definition par_cp :: "('s,'p,'f,'e) body ⇒ ('s,'p,'f,'e) com list ⇒ ('s,'f) xstate ⇒ (('s,'p,'f,'e) par_confs) set" 
where
  "par_cp Γ P s ≡ {(Γ1,l). l!0=(P,s) ∧ (Γ,l) ∈ par_cptn ∧ Γ1=Γ}" *)

lemma par_cptn_dest:"(Γ,(P,s)#(Q,t)#xs) ∈ par_cptn ⟹ (Γ,(Q,t)#xs)∈ par_cptn"
by (auto dest: par_cptn_elim_cases)


text {*
lemmas about single step computation
*}


(* ************************************************************************ *)
subsection {* Structural Properties of Small Step Computations *}
(* ************************************************************************ *)
lemma redex_not_Seq: "redex c = Seq c1 c2 ⟹ P"
  apply (induct c)
  apply auto
  done
lemma redex_not_Catch: "redex c = Catch c1 c2 ⟹ P"
  apply (induct c)
  apply auto
  done

lemma no_step_final: 
  assumes step: "Γ⊢c(c,s) → (c',s')"
  shows "final (c,s) ⟹ P"
using step
by induct (auto simp add: final_def)


lemma no_step_final':
  assumes step: "Γ⊢c cfg → cfg'"
  shows "final cfg ⟹ P"
using step
  by (cases cfg, cases cfg') (auto intro: no_step_final)


lemma step_Abrupt:
  assumes step: "Γ⊢c (c, s) → (c', s')"
  shows "⋀x. s=Abrupt x ⟹ s'=Abrupt x"
using step
by (induct) auto

lemma step_Fault: 
  assumes step: "Γ⊢c (c, s) → (c', s')"
  shows "⋀f. s=Fault f ⟹ s'=Fault f"
using step
by (induct) auto

lemma step_Stuck: 
  assumes step: "Γ⊢c (c, s) → (c', s')"
  shows "⋀f. s=Stuck ⟹ s'=Stuck"
using step
by (induct) auto

lemma step_not_normal_not_normal:
  assumes step:"Γ⊢c (c, s) → (c', s')"
  shows "∀s1. s≠Normal s1 ⟹ ∀s1. s' ≠ Normal s1"
using step step_Abrupt step_Stuck step_Fault
by (induct) auto

lemma step_not_normal_s_eq_t:
  assumes step:"Γ⊢c (c, s) → (c', t)"
  shows "∀s1. s≠Normal s1 ⟹ s=t"
using step step_Abrupt step_Stuck step_Fault
by (induct) auto

lemma ce_not_normal_s:
   assumes a1:"Γ⊢c cf0 →ce cf1"
   shows "⋀ c1 c2 s s'. ⟦cf0 = (c1,s);cf1=(c2,s');(∀sa. (s≠Normal sa))⟧
                     ⟹ s=s'"
using a1
apply (clarify, cases rule:step_ce.cases)
by (metis step_not_normal_s_eq_t env_not_normal_s)+

lemma SeqSteps: 
  assumes steps: "Γ⊢c cfg1* cfg2"
  shows "⋀ c1 s c1' s'. ⟦cfg1 = (c1,s);cfg2=(c1',s')⟧
          ⟹ Γ⊢c(Seq c1 c2,s) →* (Seq c1' c2, s')"
using steps
proof (induct rule: converse_rtranclp_induct [case_names Refl Trans]) 
  case Refl
  thus ?case
    by simp
next
  case (Trans cfg1 cfg'')      
  have step: "Γ⊢c cfg1 → cfg''" using Trans.hyps(1) by blast
  have steps: "Γ⊢c cfg'' →* cfg2" by fact
  have cfg1: "cfg1 = (c1, s)" and cfg2: "cfg2 = (c1', s')"  by fact+
  obtain c1'' s'' where cfg'': "cfg''=(c1'',s'')"
    by (cases cfg'') auto
  from step cfg1 cfg'' 
  have "Γ⊢c (c1,s) → (c1'',s'')"
    by simp  
  hence "Γ⊢c (Seq c1 c2,s) → (Seq c1'' c2,s'')" by (simp add: Seqc)    
  also from Trans.hyps (3) [OF cfg'' cfg2]
  have "Γ⊢c (Seq c1'' c2, s'') →* (Seq c1' c2, s')" .
  finally show ?case .
qed

lemma CatchSteps: 
  assumes steps: "Γ⊢ccfg1* cfg2"
  shows "⋀ c1 s c1' s'. ⟦cfg1 = (c1,s); cfg2=(c1',s')⟧
          ⟹ Γ⊢c(Catch c1 c2,s) →* (Catch c1' c2, s')"
using steps
proof (induct rule: converse_rtranclp_induct [case_names Refl Trans])
  case Refl
  thus ?case
    by simp
next
  case (Trans cfg1 cfg'')
  have step: "Γ⊢c cfg1 → cfg''" by fact
  have steps: "Γ⊢c cfg'' →* cfg2" by fact
  have cfg1: "cfg1 = (c1, s)" and cfg2: "cfg2 = (c1', s')"  by fact+
  obtain c1'' s'' where cfg'': "cfg''=(c1'',s'')"
    by (cases cfg'') auto
  from step cfg1 cfg'' 
  have s: "Γ⊢c (c1,s) → (c1'',s'')"
    by simp
  hence "Γ⊢c (Catch c1 c2,s) → (Catch c1'' c2,s'')"
    by (rule stepc.Catchc)
  also from Trans.hyps (3) [OF cfg'' cfg2]
  have "Γ⊢c (Catch c1'' c2, s'') →* (Catch c1' c2, s')" .
  finally show ?case .      
qed

lemma steps_Fault: "Γ⊢c (c, Fault f) →* (Skip, Fault f)"
proof (induct c)
  case (Seq c1 c2)
  have steps_c1: "Γ⊢c (c1, Fault f) →* (Skip, Fault f)" by fact
  have steps_c2: "Γ⊢c (c2, Fault f) →* (Skip, Fault f)" by fact
  from SeqSteps [OF steps_c1 refl refl]
  have "Γ⊢c (Seq c1 c2, Fault f) →* (Seq Skip c2, Fault f)".
  also
  have "Γ⊢c (Seq Skip c2, Fault f) → (c2, Fault f)" by (rule SeqSkipc)
  also note steps_c2
  finally show ?case by simp
next
  case (Catch c1 c2)
  have steps_c1: "Γ⊢c (c1, Fault f) →* (Skip, Fault f)" by fact
  from CatchSteps [OF steps_c1 refl refl]
  have "Γ⊢c (Catch c1 c2, Fault f) →* (Catch Skip c2, Fault f)".
  also
  have "Γ⊢c (Catch Skip c2, Fault f) → (Skip, Fault f)" by (rule CatchSkipc) 
  finally show ?case by simp
qed (fastforce intro: stepc.intros)+


lemma steps_Stuck: "Γ⊢c (c, Stuck) →* (Skip, Stuck)"
proof (induct c)
  case (Seq c1 c2)
  have steps_c1: "Γ⊢c (c1, Stuck) →* (Skip, Stuck)" by fact
  have steps_c2: "Γ⊢c (c2, Stuck) →* (Skip, Stuck)" by fact
  from SeqSteps [OF steps_c1 refl refl]
  have "Γ⊢c (Seq c1 c2, Stuck) →* (Seq Skip c2, Stuck)".
  also
  have "Γ⊢c (Seq Skip c2, Stuck) → (c2, Stuck)" by (rule SeqSkipc)
  also note steps_c2
  finally show ?case by simp
next
  case (Catch c1 c2)
  have steps_c1: "Γ⊢c (c1, Stuck) →* (Skip, Stuck)" by fact
  from CatchSteps [OF steps_c1 refl refl]
  have "Γ⊢c (Catch c1 c2, Stuck) →* (Catch Skip c2, Stuck)" .
  also
  have "Γ⊢c (Catch Skip c2, Stuck) → (Skip, Stuck)" by (rule CatchSkipc) 
  finally show ?case by simp
qed (fastforce intro: stepc.intros)+

lemma steps_Abrupt: "Γ⊢c (c, Abrupt s) →* (Skip, Abrupt s)"
proof (induct c)
  case (Seq c1 c2)
  have steps_c1: "Γ⊢c (c1, Abrupt s) →* (Skip, Abrupt s)" by fact
  have steps_c2: "Γ⊢c (c2, Abrupt s) →* (Skip, Abrupt s)" by fact
  from SeqSteps [OF steps_c1 refl refl]
  have "Γ⊢c (Seq c1 c2, Abrupt s) →* (Seq Skip c2, Abrupt s)".
  also
  have "Γ⊢c (Seq Skip c2, Abrupt s) → (c2, Abrupt s)" by (rule SeqSkipc)
  also note steps_c2
  finally show ?case by simp
next
  case (Catch c1 c2)
  have steps_c1: "Γ⊢c (c1, Abrupt s) →* (Skip, Abrupt s)" by fact
  from CatchSteps [OF steps_c1 refl refl]
  have "Γ⊢c (Catch c1 c2, Abrupt s) →* (Catch Skip c2, Abrupt s)".
  also
  have "Γ⊢c (Catch Skip c2, Abrupt s) → (Skip, Abrupt s)" by (rule CatchSkipc) 
  finally show ?case by simp
qed (fastforce intro: stepc.intros)+

lemma step_Fault_prop: 
  assumes step: "Γ⊢c (c, s) → (c', s')"
  shows "⋀f. s=Fault f ⟹ s'=Fault f"
using step
by (induct) auto

lemma step_Abrupt_prop: 
  assumes step: "Γ⊢c (c, s) → (c', s')"
  shows "⋀x. s=Abrupt x ⟹ s'=Abrupt x"
using step
by (induct) auto

lemma step_Stuck_prop: 
  assumes step: "Γ⊢c (c, s) → (c', s')"
  shows "s=Stuck ⟹ s'=Stuck"
using step
by (induct) auto

lemma steps_Fault_prop: 
  assumes step: "Γ⊢c (c, s) →* (c', s')"
  shows "s=Fault f ⟹ s'=Fault f"
using step
proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans])
  case Refl thus ?case by simp
next
  case (Trans c s c'' s'')
  thus ?case by (simp add: step_Fault_prop)    
qed

lemma steps_Abrupt_prop: 
  assumes step: "Γ⊢c (c, s) →* (c', s')"
  shows "s=Abrupt t ⟹ s'=Abrupt t"
using step
proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans])
  case Refl thus ?case by simp
next
  case (Trans c s c'' s'')
  thus ?case
    by (auto intro: step_Abrupt_prop)
qed

lemma steps_Stuck_prop: 
  assumes step: "Γ⊢c (c, s) →* (c', s')"
  shows "s=Stuck ⟹ s'=Stuck"
using step
proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans])
  case Refl thus ?case by simp   
next
  case (Trans c s c'' s'')
  thus ?case
    by (auto intro: step_Stuck_prop)
qed

lemma step_seq_throw_normal:
assumes step: "Γ⊢c (c, s) → (c', s')" and
        c_val: "c=Seq Throw Q ∧ c'=Throw"
shows "∃sa. s=Normal sa"
using step c_val
proof (cases s)
  case Normal 
  thus "∃sa. s=Normal sa" by auto
next
  case Abrupt
  thus "∃sa. s=Normal sa" using step c_val stepc_elim_cases(5)[of Γ Throw Q s "(Throw,s')"] by auto 
next 
  case Stuck
  thus "∃sa. s=Normal sa" using step c_val stepc_elim_cases(5)[of Γ Throw Q s "(Throw,s')"] by auto
next
  case Fault
    thus "∃sa. s=Normal sa" using step c_val stepc_elim_cases(5)[of Γ Throw Q s "(Throw,s')"] by auto
qed


lemma step_catch_throw_normal:
assumes step: "Γ⊢c (c, s) → (c', s')" and
        c_val: "c=Catch Throw Q ∧ c'=Throw"
shows "∃sa. s=Normal sa"
using step c_val
proof (cases s)
  case Normal 
  thus "∃sa. s=Normal sa" by auto
next
  case Abrupt
  thus "∃sa. s=Normal sa" using step c_val stepc_elim_cases(12)[of Γ Throw Q s "(Throw,s')"] by auto 
next 
  case Stuck
  thus "∃sa. s=Normal sa" using step c_val stepc_elim_cases(12)[of Γ Throw Q s "(Throw,s')"] by auto
next
  case Fault
    thus "∃sa. s=Normal sa" using step c_val stepc_elim_cases(12)[of Γ Throw Q s "(Throw,s')"] by auto
qed

lemma step_normal_to_normal[rule_format]:
assumes step:"Γ⊢c (c, s) →* (c', s')" and
        sn: "s = Normal sa" and
        finalc':"(Γ⊢c (c', s') →*(c1, s1) ∧  (∃sb. s1 = Normal sb))"
shows " (∃sc. s'=Normal sc)"
using step sn finalc'                                  
 proof (induct arbitrary: sa rule: converse_rtranclp_induct2 [case_names Refl Trans])
   case Refl show ?case by (simp add: Refl.prems)              
 next     
   case (Trans c s c'' s'') thm converse_rtranclpE2 
     thus ?case
     proof (cases s'')  
      case (Abrupt a1) thus ?thesis using finalc' by (metis steps_Abrupt_prop Trans.hyps(2))                
     next
      case Stuck thus ?thesis using finalc' by (metis steps_Stuck_prop Trans.hyps(2))          
     next
      case Fault thus ?thesis using finalc' by (metis steps_Fault_prop Trans.hyps(2))        
     next 
      case Normal thus ?thesis using Trans.hyps(3) finalc' by blast 
    qed 
qed

lemma step_spec_skip_normal_normal:
  assumes a0:"Γ⊢c (c,s)  → (c',s')" and
          a1:"c=Spec r e" and
          a2: "s=Normal s1" and
          a3: "c'=Skip" and
          a4: "(∃t. (s1,t) ∈ r)"
  shows "∃s1'. s'=Normal s1'"
proof (cases s')  
  case (Normal u) thus ?thesis by auto
next
  case Stuck 
    have "∀f r b p e. ¬ f⊢c (LanguageCon.com.Spec r e, Normal b) → p ∨ 
            (∃ba. p = (Skip::('b, 'a, 'c,'d) com, Normal ba) ∧ (b, ba) ∈ r) ∨ 
            p = (Skip, Stuck) ∧ (∀ba. (b, ba) ∉ r)"
      by (meson stepc_Normal_elim_cases(4))
      thus ?thesis using a0 a1 a2 a4 by blast
next
  case (Fault f) 
  have "∀f r b p e. ¬ f⊢c (LanguageCon.com.Spec r e, Normal b) → p ∨ 
            (∃ba. p = (Skip::('b, 'a, 'c,'d) com, Normal ba) ∧ (b, ba) ∈ r) ∨ 
            p = (Skip, Stuck) ∧ (∀ba. (b, ba) ∉ r)"
    by (meson stepc_Normal_elim_cases(4))
    thus ?thesis using a0 a1 a2 a4 by blast                       
next
  have "∀f r b p e. ¬ f⊢c (LanguageCon.com.Spec r e, Normal b) → p ∨ 
        (∃ba. p = (Skip::('b, 'a, 'c,'d) com, Normal ba) ∧ (b, ba) ∈ r) ∨ 
        p = (Skip, Stuck) ∧ (∀ba. (b, ba) ∉ r)"
    by (meson stepc_Normal_elim_cases(4))
    thus ?thesis using a0 a1 a2 a4 by blast
qed


text{* if not Normal not environmental *}

(* 

NOTE
Call always change the program now 

lemma no_advance_call_inf:
assumes a0: "redex p1 = Call f" and
        a1: "(Γ f) = Some (Call f)" 
shows "Γ⊢c (p1,Normal s) → (p1, Normal s)"
using a0 a1
proof (induct p1)
  case (Catch) thus ?case by (simp add: Catchc)
next
  case (Seq) thus ?case by (simp add: Seqc)
next
  case (Call) thus ?case
    by (simp add: Callc) 
qed(auto) *)

lemma no_advance_seq:
assumes a0: "P = Seq p1 p2" and
        a1: "Γ⊢c (p1,Normal s) → (p1, Normal s)"
shows "Γ⊢c (P,Normal s) → (P, Normal s)"
by (simp add: Seqc a0 a1)

lemma no_advance_catch:
assumes a0: "P = Catch p1 p2" and
        a1: "Γ⊢c (p1,Normal s) → (p1, Normal s)"
shows "Γ⊢c (P,Normal s) → (P, Normal s)"
by (simp add: Catchc a0 a1)

lemma not_step_c_env: 
"Γ⊢c (c, s) →e (c, s') ⟹ 
 (⋀sa. ¬(s=Normal sa)) ⟹ 
  (⋀sa. ¬(s'=Normal sa))"
by (fastforce elim:stepe_elim_cases)

lemma step_c_env_not_normal_eq_state: 
"Γ⊢c (c, s) →e (c, s') ⟹ 
 (⋀sa. ¬(s=Normal sa)) ⟹ 
  s=s'"
by (fastforce elim:stepe_elim_cases)

lemma not_eq_not_env:
   assumes step_m: "Γ⊢c (c, s) →ce (c', s')" 
   shows "~(c=c') ⟹ Γ⊢c (c, s) →e (c', s') ⟹ P"
using step_m etranE by blast


lemma step_ce_not_step_e_step_c:
   assumes step_m: "Γ⊢c (c, s) →ce (c', s')" 
   shows "¬ (Γ⊢c (c, s) →e (c', s')) ⟹(Γ⊢c (c, s) → (c', s'))"
using step_m  step_ce_elim_cases by blast

lemma step_ce_notNormal:   
   assumes step_m: "Γ⊢c (c, s) →ce (c', s')" 
   shows "(∀sa. ¬(s=Normal sa)) ⟹ s'=s"
using step_m
proof (induct rule:step_ce_induct) 
  case (e_step a b a' b')
  have "∀f p pa. ¬ f⊢c p →e pa ∨ (∃c. (∃x. p = (c::('b, 'a, 'c,'d) LanguageCon.com, x)) ∧ (∃x. pa = (c, x)))"
    by (fastforce elim:etranE stepe_elim_cases)
  thus ?case
    using stepe_elim_cases e_step.hyps e_step.prems by blast
next
  case (c_step a b a' b')
  thus ?case 
  proof (cases b)
    case (Normal) thus ?thesis using c_step.prems by auto   
  next
    case (Stuck) thus ?thesis  
      using SmallStepCon.step_Stuck_prop c_step.hyps by blast 
  next
    case (Fault f) thus ?thesis 
     using SmallStepCon.step_Fault_prop c_step.hyps by fastforce
  next
    case (Abrupt a) thus ?thesis
      using SmallStepCon.step_Abrupt_prop c_step.hyps by fastforce 
  qed
qed

lemma steps_ce_not_Normal:  
   assumes step_m: "Γ⊢c (c, s) →ce* (c', s')" 
   shows "∀sa. ¬(s=Normal sa) ⟹ s'=s"
using step_m
proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans])
  case Refl then show ?case by auto
next
  case (Trans a b a' b') 
  thus ?case using step_ce_notNormal by blast 
qed

lemma steps_not_normal_ce_c: 
  assumes steps: "Γ⊢c (c, s) →ce* (c', s')"
  shows         "( ∀sa. ¬(s=Normal sa)) ⟹ Γ⊢c (c, s) →* (c', s')"
using steps
proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans])
  case Refl thus ?case by auto
next
  case (Trans a b a' b') 
    then have "b=b'" using step_ce_notNormal by blast
    then have "Γ⊢c (a', b') →* (c', s')" using `b=b'` Trans.hyps(3) Trans.prems by blast
    then have "Γ⊢c (a, b) → (a', b') ∨ Γ⊢c (a, b) →e (a', b')"
      using Trans.hyps(1) by (fastforce elim: step_ce_elim_cases)
    thus ?case
    proof
      assume "Γ⊢c (a, b) → (a', b')" 
      thus ?thesis using `Γ⊢c (a', b') →* (c', s')` by auto 
    next
      assume "Γ⊢c (a, b) →e (a', b')"   
       have "a = a'"
         by (meson Trans.hyps(1) `Γ⊢c (a, b) →e (a', b')` not_eq_not_env)   
         thus ?thesis using `Γ⊢c (a', b') →* (c', s')` `b = b'` by force           
    qed
qed

lemma steps_c_ce: 
  assumes steps: "Γ⊢c (c, s) →* (c', s')"
  shows         "Γ⊢c (c, s) →ce* (c', s')"
using steps
proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans])
  case Refl thus ?case by auto
next
  case (Trans a b a' b') 
  have "Γ⊢c (a, b) →ce (a', b')"
    using Trans.hyps(1) c_step by blast
  thus ?case
    by (simp add: Trans.hyps(3) converse_rtranclp_into_rtranclp)
qed

lemma steps_not_normal_c_ce: 
  assumes steps: "Γ⊢c (c, s) →* (c', s')"
  shows         "( ∀sa. ¬(s=Normal sa)) ⟹ Γ⊢c (c, s) →ce* (c', s')"
by (simp add: steps steps_c_ce)

lemma steps_not_normal_c_eq_ce:
assumes normal: "( ∀sa. ¬(s=Normal sa))"
shows         " Γ⊢c (c, s) →* (c', s') =  Γ⊢c (c, s) →ce* (c', s')"
using normal
using steps_c_ce steps_not_normal_ce_c by auto 

lemma steps_ce_Fault: "Γ⊢c (c, Fault f) →ce* (Skip, Fault f)"
by (simp add: SmallStepCon.steps_Fault steps_c_ce)

lemma steps_ce_Stuck: "Γ⊢c (c, Stuck) →ce* (Skip, Stuck)"
by (simp add: SmallStepCon.steps_Stuck steps_c_ce)

lemma steps_ce_Abrupt: "Γ⊢c (c, Abrupt a) →ce* (Skip, Abrupt a)"
by (simp add: SmallStepCon.steps_Abrupt steps_c_ce)

lemma step_ce_seq_throw_normal:
assumes step: "Γ⊢c (c, s) →ce (c', s')" and
        c_val: "c=Seq Throw Q ∧ c'=Throw"
shows "∃sa. s=Normal sa"
using step c_val not_eq_not_env 
      step_ce_not_step_e_step_c step_seq_throw_normal by blast

lemma step_ce_catch_throw_normal:
assumes step: "Γ⊢c (c, s) →ce (c', s')" and
        c_val: "c=Catch Throw Q ∧ c'=Throw"
shows "∃sa. s=Normal sa"
using step c_val not_eq_not_env
      step_ce_not_step_e_step_c step_catch_throw_normal  by blast 

lemma step_ce_normal_to_normal[rule_format]:
assumes step:"Γ⊢c (c, s) →ce* (c', s')" and
        sn: "s = Normal sa" and
        finalc':"(Γ⊢c (c', s') →ce*(c1, s1) ∧  (∃sb. s1 = Normal sb))"
shows "         
       (∃sc. s'=Normal sc)"
using step sn finalc' steps_ce_not_Normal by blast      

lemma SeqSteps_ce: 
  assumes steps: "Γ⊢c cfg1ce* cfg2"
  shows "⋀ c1 s c1' s'. ⟦cfg1 = (c1,s);cfg2=(c1',s')⟧
          ⟹ Γ⊢c(Seq c1 c2,s) →ce* (Seq c1' c2, s')"
using steps
proof (induct rule: converse_rtranclp_induct [case_names Refl Trans]) 
  case Refl
  thus ?case
    by simp
next
  case (Trans cfg1 cfg'') 
  then have "Γ⊢c cfg1 → cfg'' ∨ Γ⊢c cfg1e cfg''" 
   using step_ce_elim_cases by blast
  thus ?case 
  proof
    assume a1:"Γ⊢c cfg1e cfg''"          
    have "∀f p pa. ¬ f⊢c p →e pa ∨ (∃c. 
                   (∃x. p = (c::('a, 'b, 'c,'d) LanguageCon.com, x)) ∧ (∃x. pa = (c, x)))"
      by (meson etranE)
    then obtain cc :: "('b ⇒ ('a, 'b, 'c,'d) LanguageCon.com option) ⇒ 
                              ('a, 'b, 'c,'d) LanguageCon.com × ('a, 'c) xstate ⇒ 
                              ('a, 'b, 'c,'d) LanguageCon.com × ('a, 'c) xstate ⇒ 
                              ('a, 'b, 'c,'d) LanguageCon.com" and 
                xx :: "('b ⇒ ('a, 'b, 'c,'d) LanguageCon.com option) ⇒ 
                       ('a, 'b, 'c,'d) LanguageCon.com × ('a, 'c) xstate ⇒ 
                       ('a, 'b, 'c,'d) LanguageCon.com × ('a, 'c) xstate ⇒ ('a, 'c) xstate" and
                xxa :: "('b ⇒ ('a, 'b, 'c,'d) LanguageCon.com option) ⇒ 
                        ('a, 'b, 'c,'d) LanguageCon.com × ('a, 'c) xstate ⇒ 
                       ('a, 'b, 'c,'d) LanguageCon.com × ('a, 'c) xstate ⇒ ('a, 'c) xstate" where
      f1: "∀f p pa. ¬ f⊢c p →e pa ∨ p = (cc f p pa, xx f p pa) ∧ pa = (cc f p pa, xxa f p pa)"
      by (metis (no_types))
    have f2: "∀f c x xa. ¬ f⊢c (c::('a, 'b, 'c,'d) LanguageCon.com, x) →e (c, xa) ∨ 
                            (∃a. x = Normal a) ∨ (∀a. xa ≠ Normal a) ∧ x = xa"
      by (metis stepe_elim_cases)
    have f3: "(c1, xxa Γ cfg1 cfg'') = cfg''"
      using f1 by (metis Trans.prems(1) a1 fst_conv)
    hence "Γ⊢c (LanguageCon.com.Seq c1 c2, xxa Γ cfg1 cfg'') →ce* (LanguageCon.com.Seq c1' c2, s')"
      using Trans.hyps(3) Trans.prems(2) by force
    thus ?thesis
      using f3 f2 by (metis (no_types) Env Trans.prems(1) a1 e_step r_into_rtranclp 
                       rtranclp.rtrancl_into_rtrancl rtranclp_idemp)   
  next
     assume "Γ⊢c cfg1 → cfg''"
     thus ?thesis
      proof -
        have "∀p. ∃c x. p = (c::('a, 'b, 'c,'d) LanguageCon.com, x::('a, 'c) xstate)"
          by auto
        thus ?thesis
          by (metis (no_types) Seqc Trans.hyps(3) Trans.prems(1) Trans.prems(2) 
                   `Γ⊢c cfg1 → cfg''` c_step converse_rtranclp_into_rtranclp)
      qed
  qed
qed


lemma CatchSteps_ce: 
  assumes steps: "Γ⊢ccfg1ce* cfg2"
  shows "⋀ c1 s c1' s'. ⟦cfg1 = (c1,s); cfg2=(c1',s')⟧
          ⟹ Γ⊢c(Catch c1 c2,s) →ce* (Catch c1' c2, s')"
using steps
proof (induct rule: converse_rtranclp_induct [case_names Refl Trans])
  case Refl
  thus ?case
    by simp
next
  case (Trans cfg1 cfg'')
then have "Γ⊢c cfg1 → cfg'' ∨ Γ⊢c cfg1e cfg''" 
   using step_ce_elim_cases by blast
  thus ?case 
  proof
    assume a1:"Γ⊢c cfg1e cfg''"        
    have "∀f p pa. ¬ f⊢c p →e pa ∨ (∃c. (∃x. p = (c::('a, 'b, 'c,'d) LanguageCon.com, x)) ∧ (∃x. pa = (c, x)))"
      by (meson etranE)
    then obtain cc :: "('b ⇒('a, 'b, 'c, 'd) LanguageCon.com option) ⇒
                        ('a, 'b, 'c, 'd) LanguageCon.com × ('a, 'c) xstate ⇒
                        ('a, 'b, 'c, 'd) LanguageCon.com × ('a, 'c) xstate ⇒
                        ('a, 'b, 'c, 'd) LanguageCon.com" and 
                xx :: "('b ⇒('a, 'b, 'c, 'd) LanguageCon.com option) ⇒
                       ('a, 'b, 'c, 'd) LanguageCon.com × ('a, 'c) xstate ⇒
                       ('a, 'b, 'c, 'd) LanguageCon.com × ('a, 'c) xstate ⇒ 
                       ('a, 'c) xstate" and 
                xxa :: "('b ⇒('a, 'b, 'c, 'd) LanguageCon.com option) ⇒
                         ('a, 'b, 'c, 'd) LanguageCon.com × ('a, 'c) xstate ⇒
                         ('a, 'b, 'c, 'd) LanguageCon.com × ('a, 'c) xstate ⇒ ('a, 'c) xstate" where
      f1: "∀f p pa. ¬ f⊢c p →e pa ∨ p = (cc f p pa, xx f p pa) ∧ pa = (cc f p pa, xxa f p pa)"
      by (metis (no_types))
    have f2: "∀f c x xa. ¬ f⊢c (c::('a, 'b, 'c,'d) LanguageCon.com, x) →e (c, xa) ∨ 
                         (∃a. x = Normal a) ∨ (∀a. xa ≠ Normal a) ∧ x = xa"
      by (metis stepe_elim_cases)
    have f3: "(c1, xxa Γ cfg1 cfg'') = cfg''"
      using f1 by (metis Trans.prems(1) a1 fst_conv)
    hence "Γ⊢c (LanguageCon.com.Catch c1 c2, xxa Γ cfg1 cfg'') →ce* (LanguageCon.com.Catch c1' c2, s')"
      using Trans.hyps(3) Trans.prems(2) by force
    thus ?thesis
      using f3 f2 by (metis (no_types) Env Trans.prems(1) a1 e_step r_into_rtranclp rtranclp.rtrancl_into_rtrancl rtranclp_idemp)              
  next
    assume "Γ⊢c cfg1 → cfg''"
    thus ?thesis
    proof -
      obtain cc :: "('a, 'b, 'c, 'd) LanguageCon.com × ('a, 'c) xstate ⇒ ('a, 'b, 'c, 'd) LanguageCon.com" and xx :: "('a, 'b, 'c, 'd) LanguageCon.com × ('a, 'c) xstate ⇒ ('a, 'c) xstate" where
        f1: "∀p. p = (cc p, xx p)"
        by (meson old.prod.exhaust)
      hence "⋀c. Γ⊢c (LanguageCon.com.Catch c1 c, s) → (LanguageCon.com.Catch (cc cfg'') c, xx cfg'')"
        by (metis (no_types) Catchc Trans.prems(1) `Γ⊢c cfg1 → cfg''`)
      thus ?thesis
        using f1 by (meson Trans.hyps(3) Trans.prems(2) c_step converse_rtranclp_into_rtranclp)
    qed      
  qed
qed

lemma step_change_p_or_eq_Ns: 
    assumes step: "Γ⊢c (P,Normal s) → (Q,s')"
    shows  "¬(P=Q)"
using step
proof (induct P arbitrary: Q s s')
qed(fastforce elim: stepc_Normal_elim_cases)+

 
lemma step_change_p_or_eq_s: 
    assumes step: "Γ⊢c (P,s) → (Q,s')"
    shows  "¬(P=Q)"
using step
proof (induct P arbitrary: Q s s')
qed (fastforce elim: stepc_elim_cases)+

subsection {* Relation between @{term "stepc_rtrancl"} and @{term "cptn"} *}

lemma stepc_rtrancl_cptn:
  assumes step: "Γ⊢c (c,s) →ce* (cf,sf)"           
  shows "∃xs. (Γ,(c, s)#xs) ∈ cptn ∧(cf,sf) = (last ((c,s)#xs))"
using step 
proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans]) 
  case Refl thus ?case using cptn.CptnOne by auto
next
  case (Trans c s c' s')  
  have "Γ⊢c (c, s) →e (c', s') ∨ Γ⊢c (c, s) → (c', s')"
    by (meson Trans.hyps(1) step_ce.simps)
  then show ?case
  proof
    assume prem:"Γ⊢c (c, s) →e (c', s')"       
    then have ceqc':"c=c'" using prem env_c_c'
      by auto           
    obtain xs where xs_s:"(Γ, (c', s') # xs) ∈ cptn ∧ (cf, sf) = last ((c', s') # xs)"
      using Trans(3) by auto 
    then have xs_f: "(Γ, (c, s)#(c', s') # xs) ∈ cptn" 
    using cptn.CptnEnv ceqc'  prem by fastforce
    also have "last ((c', s') # xs) = last ((c,s)#(c', s') # xs)" by auto
    then have "(cf, sf) = last ((c, s) # (c', s') # xs)"
      using   xs_s by auto
    thus ?thesis
      using  xs_f by blast
  next
    assume prem:"Γ⊢c (c, s) → (c', s')" 
    obtain xs where xs_s:"(Γ, (c', s') # xs) ∈ cptn ∧ (cf, sf) = last ((c', s') # xs) "
      using Trans(3) by auto 
    have "(Γ, (c, s) # (c', s') # xs) ∈ cptn" using cptn.CptnComp 
      using xs_s prem by blast 
    also have "last ((c', s') # xs) = last ((c,s)#(c', s') # xs)" by auto
    ultimately show ?thesis using xs_s by fastforce
  qed 
qed


lemma cptn_stepc_rtrancl:
  assumes cptn_step: "(Γ,(c, s)#xs) ∈ cptn" and
          cf_last:"(cf,sf) = (last ((c,s)#xs))"
  shows "Γ⊢c (c,s) →ce* (cf,sf)"
using cptn_step cf_last
proof (induct xs arbitrary: c s) 
  case Nil
  thus ?case by simp 
next
  case (Cons a xs c s) 
  then obtain ca sa where eq_pair: "a=(ca,sa)" and "(cf, sf) = last ((ca,sa) # xs) " 
       using Cons by (fastforce)  
  have f1: "∀f p pa. ¬ (f::'a ⇒ ('b, _, 'c,'d) LanguageCon.com option)⊢c p → pa ∨ f⊢c p →ce pa"
    by (simp add: c_step)
  have f2: "(Γ, (c, s) # (ca, sa) # xs) ∈ cptn"
    using `(Γ, (c, s) # a # xs) ∈ cptn` eq_pair by blast
  have f3: "∀f p pa. ¬ (f::'a ⇒ ('b, _, 'c,'d) LanguageCon.com option)⊢c p →e pa ∨ f⊢c p →ce pa"
    using e_step by blast
  have "∀c x. (Γ, (c, x) # xs) ∉ cptn ∨ (cf, sf) ≠ last ((c, x) # xs) ∨ Γ⊢c (c, x) →ce* (cf, sf)"
    using Cons.hyps by blast
  thus ?case
    using f3 f2 f1 by (metis (no_types) `(cf, sf) = last ((ca, sa) # xs)` converse_rtranclp_into_rtranclp cptn_elim_cases(2)) 
qed

lemma three_elems_list:
  assumes a1:"length l > 2"
  shows "∃a0 a1 a2 l1. l=a0#a1#a2#l1"
using a1 by (metis Cons_nth_drop_Suc One_nat_def Suc_1 Suc_leI add_lessD1 drop_0 length_greater_0_conv list.size(3) not_numeral_le_zero one_add_one)  

lemma cptn_stepc_rtran:
  assumes cptn_step: "(Γ,x#xs) ∈ cptn" and          
          a1:"Suc i < length (x#xs)"
  shows "Γ⊢c ((x#xs)!i) →ce ((x#xs)!(Suc i))"
using cptn_step a1
proof (induct i arbitrary: x xs)
  case 0
    then obtain x1 xs1 where xs:"xs=x1#xs1"
      by (metis length_Cons less_not_refl list.exhaust list.size(3))    
    then have "(x#x1#xs1)!Suc 0 = x1" by fastforce  
    have x_x1_cptn:"(Γ,x#x1#xs1)∈cptn" using 0 xs by auto    
    then have "(Γ,x1#xs1)∈cptn"
      using cptn_dest_pair by fastforce
    then have "Γ⊢cx →e x1 ∨ Γ⊢cx → x1" 
      using cptn_elim_cases_pair x_x1_cptn  by blast        
    then have "Γ⊢c x →ce x1"
      by (metis c_step e_step)
    then show ?case 
      by (simp add: xs)
next
   case (Suc i)
    then have "Suc i < length xs" by auto
    moreover then obtain x1 xs1 where xs:"xs=x1#xs1"
      by (metis (full_types) list.exhaust list.size(3) not_less0)
    moreover then have "(Γ,x1#xs1) ∈ cptn" using Suc cptn_dest_pair by blast
    ultimately have "Γ⊢c ((x1 # xs1) ! i) →ce ((x1 # xs1) ! Suc i)" 
      using Suc by auto
    thus ?case using Suc xs by auto
qed 
     
      
lemma cptn_stepconf_rtrancl:
  assumes cptn_step: "(Γ,cfg1#xs) ∈ cptn" and
          cf_last:"cfg2 = (last (cfg1#xs))"
  shows "Γ⊢c cfg1 →ce* cfg2"
using cptn_step cf_last
by (metis cptn_stepc_rtrancl prod.collapse)

lemma cptn_all_steps_rtrancl:
  assumes cptn_step: "(Γ,cfg1#xs) ∈ cptn"          
  shows "∀i<length (cfg1#xs). Γ⊢c cfg1 →ce* ((cfg1#xs)!i)"
using cptn_step 
proof (induct xs arbitrary: cfg1)
  case Nil thus ?case by auto
next
  case (Cons x xs1) thus ?case
  proof -
     have hyp:"∀i<length (x # xs1). Γ⊢c x →ce* ((x # xs1) ! i)"
       using Cons.hyps Cons.prems cptn_dest_pair by blast      
     thus ?thesis
     proof
     {
        fix i
        assume a0:"i<length (cfg1 # x # xs1)"
        then have "Suc 0 < length (cfg1 # x # xs1)"
          by simp
        hence "Γ⊢c (cfg1 # x # xs1) ! 0 →ce ((cfg1 # x # xs1) ! Suc 0)"
          using Cons.prems cptn_stepc_rtran by blast
        then have "Γ⊢c cfg1 →ce x" using Cons by simp
        also have  "i < Suc (Suc (length xs1))"
          using a0 by force
        ultimately have "Γ⊢c cfg1 →ce* (cfg1 # x # xs1) ! i" using hyp Cons
         using converse_rtranclp_into_rtranclp hyp less_Suc_eq_0_disj 
         by auto 
     } thus ?thesis by auto qed
  qed
qed

lemma cptn_env_same_prog:
assumes a0: "(Γ, l) ∈ cptn" and
        a1:  "∀k < j. (Γ⊢c(l!k)  →e (l!(Suc k)))" and
        a2: "Suc j < length l"
shows "fst (l!j) =  fst (l!0)"
using a0 a1 a2
proof (induct j arbitrary: l)
  case 0 thus ?case by auto
next
  case (Suc j) 
    then have "fst (l!j) =  fst (l!0)" by fastforce
    thus ?case using Suc 
      by (metis (no_types) env_c_c' lessI prod.collapse)
qed

lemma takecptn_is_cptn [rule_format, elim!]:
  "∀j. (Γ,c) ∈ cptn ⟶ (Γ, take (Suc j) c) ∈ cptn"
apply(induct "c")
 apply(force elim: cptn.cases)
apply clarify
apply(case_tac j)
 apply simp
 apply(rule CptnOne)
apply simp
apply(force intro:cptn.intros elim:cptn.cases)
done



lemma dropcptn_is_cptn [rule_format,elim!]:
  "∀j<length c. (Γ,c) ∈ cptn ⟶ (Γ, drop j c) ∈ cptn"
apply(induct "c")
 apply(force elim: cptn.cases)
apply clarify
apply(case_tac j,simp+)
apply(erule cptn.cases)
  apply simp
 apply force
apply force
done

lemma takepar_cptn_is_par_cptn [rule_format,elim]:
  "∀j. (Γ,c) ∈ par_cptn ⟶ (Γ,take (Suc j) c) ∈ par_cptn"
apply(induct "c")
 apply(force elim: cptn.cases)
apply clarify
apply(case_tac j,simp)
 apply(rule ParCptnOne)
apply(force intro:par_cptn.intros elim:par_cptn.cases)
done

lemma droppar_cptn_is_par_cptn [rule_format]:
  "∀j<length c. (Γ,c) ∈ par_cptn ⟶ (Γ,drop j c) ∈ par_cptn"
apply(induct "c")
 apply(force elim: par_cptn.cases)
apply clarify
apply(case_tac j,simp+)
apply(erule par_cptn.cases)
  apply simp
 apply force
apply force
done


subsection‹Modular Definition of Computation›
definition lift :: "('s,'p,'f,'e) com ⇒ ('s,'p,'f,'e) config ⇒ ('s,'p,'f,'e) config" where
  "lift Q ≡ λ(P, s).  ((Seq P Q), s)" 

definition lift_catch :: "('s,'p,'f,'e) com ⇒ ('s,'p,'f,'e) config ⇒ ('s,'p,'f,'e) config" where
  "lift_catch Q ≡ λ(P, s).  (Catch P Q, s)"     


inductive_set cptn_mod :: "(('s,'p,'f,'e) confs) set"
where
  CptnModOne: "(Γ,[(P, s)]) ∈ cptn_mod"
| CptnModEnv: "⟦Γ⊢c(P,s) →e (P,t);(Γ,(P, t)#xs) ∈ cptn_mod ⟧ ⟹ 
               (Γ,(P, s)#(P, t)#xs) ∈ cptn_mod"
| CptnModSkip: "⟦Γ⊢c(P,s) → (Skip,t); redex P = P;
                (Γ,(Skip, t)#xs) ∈ cptn_mod ⟧ ⟹ 
                (Γ,(P,s)#(Skip, t)#xs) ∈cptn_mod"

| CptnModThrow: "⟦Γ⊢c(P,s) → (Throw,t); redex P = P;
                (Γ,(Throw, t)#xs) ∈ cptn_mod ⟧ ⟹ 
                (Γ,(P,s)#(Throw, t)#xs) ∈cptn_mod"

| CptnModCondT: "⟦(Γ,(P0, Normal s)#ys) ∈ cptn_mod; s ∈ b ⟧ ⟹ 
                (Γ,((Cond b P0 P1), Normal s)#(P0, Normal s)#ys) ∈ cptn_mod"
| CptnModCondF: "⟦(Γ,(P1, Normal s)#ys) ∈ cptn_mod; s ∉ b ⟧ ⟹ 
                (Γ,((Cond b P0 P1), Normal s)#(P1, Normal s)#ys) ∈ cptn_mod"
| CptnModSeq1: 
  "⟦(Γ,(P0, s)#xs) ∈ cptn_mod; zs=map (lift P1) xs ⟧ ⟹ 
   (Γ,((Seq P0 P1), s)#zs) ∈ cptn_mod"

| CptnModSeq2: 
  "⟦(Γ, (P0, s)#xs) ∈ cptn_mod; fst(last ((P0, s)#xs)) = Skip;
    (Γ,(P1, snd(last ((P0, s)#xs)))#ys) ∈ cptn_mod;
    zs=(map (lift P1) xs)@((P1, snd(last ((P0, s)#xs)))#ys) ⟧ ⟹ 
   (Γ,((Seq P0 P1), s)#zs) ∈ cptn_mod"
(*| CptnModSeq3:
  "⟦ (Γ,(P1, s)#xs) ∈ cptn_mod⟧ ⟹ (Γ,((Seq Skip P1), s)#(P1, s)#xs) ∈ cptn_mod"*)

| CptnModSeq3: 
  "⟦(Γ, (P0, Normal s)#xs) ∈ cptn_mod; 
    fst(last ((P0, Normal s)#xs)) = Throw;
    snd(last ((P0, Normal s)#xs)) = Normal s'; 
    (Γ,(Throw,Normal s')#ys) ∈ cptn_mod; 
    zs=(map (lift P1) xs)@((Throw,Normal s')#ys) ⟧ ⟹
   (Γ,((Seq P0 P1), Normal s)#zs) ∈ cptn_mod"

| CptnModWhile1: 
  "⟦(Γ, (P, Normal s)#xs) ∈ cptn_mod; s ∈ b; 
    zs=map (lift (While b P)) xs ⟧ ⟹ 
    (Γ, ((While b P), Normal s)#
      ((Seq P (While b P)),Normal s)#zs) ∈ cptn_mod"

| CptnModWhile2: 
  "⟦ (Γ, (P, Normal s)#xs) ∈ cptn_mod; 
     fst(last ((P, Normal s)#xs))=Skip; s ∈ b; 
     zs=(map (lift (While b P)) xs)@
      (While b P, snd(last ((P, Normal s)#xs)))#ys; 
      (Γ,(While b P, snd(last ((P, Normal s)#xs)))#ys) ∈ 
        cptn_mod⟧ ⟹ 
   (Γ,(While b P, Normal s)#
     (Seq P (While b P), Normal s)#zs) ∈ cptn_mod"

| CptnModWhile3: 
  "⟦ (Γ, (P, Normal s)#xs) ∈ cptn_mod; 
     fst(last ((P, Normal s)#xs))=Throw; s ∈ b;
     snd(last ((P, Normal s)#xs)) = Normal s'; 
    (Γ,(Throw,Normal s')#ys) ∈ cptn_mod;  
     zs=(map (lift (While b P)) xs)@((Throw,Normal s')#ys)⟧ ⟹ 
   (Γ,(While b P, Normal s)#
     (Seq P (While b P), Normal s)#zs) ∈ cptn_mod"

| CptnModCall: "⟦(Γ,(bdy, Normal s)#ys) ∈ cptn_mod;Γ p = Some bdy; bdy≠Call p ⟧ ⟹ 
                (Γ,((Call p), Normal s)#(bdy, Normal s)#ys) ∈ cptn_mod" 
| CptnModDynCom: "⟦(Γ,(c s, Normal s)#ys) ∈ cptn_mod ⟧ ⟹ 
                  (Γ,(DynCom c, Normal s)#(c s, Normal s)#ys) ∈ cptn_mod"

| CptnModGuard: "⟦(Γ,(c, Normal s)#ys) ∈ cptn_mod; s ∈ g ⟧ ⟹ 
                 (Γ,(Guard f g c, Normal s)#(c, Normal s)#ys) ∈ cptn_mod"

| CptnModCatch1: "⟦(Γ,(P0, s)#xs) ∈ cptn_mod; zs=map (lift_catch P1) xs ⟧
                 ⟹ (Γ,((Catch P0 P1), s)#zs) ∈ cptn_mod"
| CptnModCatch2: 
  "⟦(Γ, (P0, s)#xs) ∈ cptn_mod; fst(last ((P0, s)#xs)) = Skip; 
    (Γ,(Skip,snd(last ((P0, s)#xs)))#ys) ∈ cptn_mod;  
    zs=(map (lift_catch P1) xs)@((Skip,snd(last ((P0, s)#xs)))#ys) ⟧ ⟹ 
   (Γ,((Catch P0 P1), s)#zs) ∈ cptn_mod"

| CptnModCatch3: 
  "⟦(Γ, (P0, Normal s)#xs) ∈ cptn_mod; fst(last ((P0, Normal s)#xs)) = Throw; 
  snd(last ((P0, Normal s)#xs)) = Normal s';
  (Γ,(P1, snd(last ((P0, Normal s)#xs)))#ys) ∈ cptn_mod; 
  zs=(map (lift_catch P1) xs)@((P1, snd(last ((P0, Normal s)#xs)))#ys) ⟧ ⟹ 
   (Γ,((Catch P0 P1), Normal s)#zs) ∈ cptn_mod"


lemmas CptnMod_induct = cptn_mod.induct [of _ "[(c,s)]", split_format (complete), case_names
CptnModOne CptnModEnv CptnModSkip CptnModThrow CptnModCondT CptnModCondF 
CptnModSeq1 CptnModSeq2 CptnModSeq3 CptnModSeq4 CptnModWhile1 CptnModWhile2 CptnModWhile3 CptnModCall CptnModDynCom CptnModGuard 
CptnModCatch1 CptnModCatch2 CptnModCatch3, induct set]

inductive_cases CptnMod_elim_cases [cases set]:
"(Γ,(Skip, s)#u#xs) ∈ cptn_mod"
"(Γ,(Guard f g c, s)#u#xs) ∈ cptn_mod"
"(Γ,(Basic f e, s)#u#xs) ∈ cptn_mod"
"(Γ,(Spec r e, s)#u#xs) ∈ cptn_mod"
"(Γ,(Seq c1 c2, s)#u#xs) ∈ cptn_mod"
"(Γ,(Cond b c1 c2, s)#u#xs) ∈ cptn_mod"
"(Γ,(Await b c2 e, s)#u#xs) ∈ cptn_mod"
"(Γ,(Call p, s)#u#xs) ∈ cptn_mod"
"(Γ,(DynCom c,s)#u#xs) ∈ cptn_mod"
"(Γ,(Throw,s)#u#xs) ∈ cptn_mod"
"(Γ,(Catch c1 c2,s)#u#xs) ∈ cptn_mod"


inductive_cases CptnMod_Normal_elim_cases [cases set]:
"(Γ,(Skip, Normal s)#u#xs) ∈ cptn_mod"
"(Γ,(Guard f g c, Normal s)#u#xs) ∈ cptn_mod"
"(Γ,(Basic f e, Normal s)#u#xs) ∈ cptn_mod"
"(Γ,(Spec r e, Normal s)#u#xs) ∈ cptn_mod"
"(Γ,(Seq c1 c2, Normal s)#u#xs) ∈ cptn_mod"
"(Γ,(Cond b c1 c2, Normal s)#u#xs) ∈ cptn_mod"
"(Γ,(Await b c2 e, Normal s)#u#xs) ∈ cptn_mod"
"(Γ,(Call p, Normal s)#u#xs) ∈ cptn_mod"
"(Γ,(DynCom c,Normal s)#u#xs) ∈ cptn_mod"
"(Γ,(Throw,Normal s)#u#xs) ∈ cptn_mod"
"(Γ,(Catch c1 c2,Normal s)#u#xs) ∈ cptn_mod"
"(Γ,(P,Normal s)#(P,s')#xs) ∈ cptn_mod"
"(Γ,(P,Abrupt s)#(P,Abrupt s')#xs) ∈ cptn_mod"
"(Γ,(P,Stuck)#(P,Stuck)#xs) ∈ cptn_mod"
"(Γ,(P,Fault f)#(P,Fault f)#xs) ∈ cptn_mod"

inductive_cases CptnMod_env_elim_cases [cases set]:
"(Γ,(P,Normal s)#(P,s')#xs) ∈ cptn_mod"
"(Γ,(P,Abrupt s)#(P,Abrupt s')#xs) ∈ cptn_mod"
"(Γ,(P,Stuck)#(P,Stuck)#xs) ∈ cptn_mod"
"(Γ,(P,Fault f)#(P,Fault f)#xs) ∈ cptn_mod"


subsection ‹Equivalence of small semantics and computational›

lemma last_length: "((a#xs)!(length xs))=last (a#xs)"
  by (induct xs) auto

definition catch_cond 
where
"catch_cond zs Q xs P s s'' s' Γ ≡ (zs=(map (lift_catch Q) xs) ∨
             ((fst(((P, s)#xs)!length xs)=Throw ∧
               snd(last ((P, s)#xs)) = Normal s' ∧ s=Normal s''∧
               (∃ys. (Γ,(Q, snd(((P, s)#xs)!length xs))#ys) ∈ cptn_mod ∧ 
                zs=(map (lift_catch Q) xs)@((Q, snd(((P, s)#xs)!length xs))#ys)))) ∨
                ((fst(((P, s)#xs)!length xs)=Skip ∧ 
                 (∃ys. (Γ,(Skip,snd(last ((P, s)#xs)))#ys) ∈ cptn_mod ∧                   
                 zs=(map (lift_catch Q) xs)@((Skip,snd(last ((P, s)#xs)))#ys)))))
"

lemma div_catch: assumes cptn_m:"(Γ,list) ∈ cptn_mod"
shows "(∀s P Q zs. list=(Catch P Q, s)#zs ⟶
       (∃xs s' s''. 
          (Γ,(P, s)#xs) ∈ cptn_mod  ∧ 
             catch_cond zs Q xs P s s'' s' Γ))  
            "
unfolding catch_cond_def
using cptn_m
proof (induct rule: cptn_mod.induct)
case (CptnModOne Γ P s)
  thus ?case using cptn_mod.CptnModOne by blast 
next
  case (CptnModSkip  Γ P s t xs) 
  from CptnModSkip.hyps
  have step: "Γ⊢c (P, s) → (Skip, t)" by auto
  from CptnModSkip.hyps
  have noskip: "~(P=Skip)" using stepc_elim_cases(1) by blast  
  have no_catch: "∀p1 p2. ¬(P=Catch p1 p2)" using CptnModSkip.hyps(2) redex_not_Catch by auto
  from CptnModSkip.hyps
  have in_cptn_mod: "(Γ, (Skip, t) # xs) ∈ cptn_mod" by auto  
  then show ?case using no_catch by simp
next
  case (CptnModThrow  Γ P s t xs) 
  from CptnModThrow.hyps
  have step: "Γ⊢c (P, s) → (Throw, t)" by auto 
  from CptnModThrow.hyps
  have in_cptn_mod: "(Γ, (Throw, t) # xs) ∈ cptn_mod" by auto 
  have no_catch: "∀p1 p2. ¬(P=Catch p1 p2)" using CptnModThrow.hyps(2) redex_not_Catch by auto
  then show ?case by auto
next
  case (CptnModCondT Γ P0 s ys b P1)
  thus ?case using CptnModOne by blast
next
  case (CptnModCondF Γ P0 s ys b P1)
  thus ?case using CptnModOne by blast
next
  case (CptnModCatch1 sa P Q zs)
  thus ?case by blast
next
  case (CptnModCatch2 Γ P0 s xs ys zs P1) 
  from CptnModCatch2.hyps(3) 
  have last:"fst (((P0, s) # xs) ! length xs) = Skip" 
       by (simp add: last_length) 
  have P0cptn:"(Γ, (P0, s) # xs) ∈ cptn_mod" by fact          
  then have "zs = map (lift_catch P1) xs @((Skip,snd(last ((P0, s)#xs)))#ys)" by (simp add:CptnModCatch2.hyps)
  show ?case
  proof -{
    fix sa P Q zsa    
    assume eq:"(Catch P0 P1, s) # zs = (Catch P Q, sa) # zsa"
    then have "P0 =P ∧ P1 = Q ∧ s=sa ∧ zs=zsa" by auto
    then have "(P0, s) = (P, sa)" by auto 
    have "last ((P0, s) # xs) = ((P, sa) # xs) ! length xs"
      by (simp add: `P0 = P ∧ P1 = Q ∧ s = sa ∧ zs = zsa` last_length)
    then have "zs = (map (lift_catch Q) xs)@((Skip,snd(last ((P0, s)#xs)))#ys)"
      using `P0 = P ∧ P1 = Q ∧ s = sa ∧ zs = zsa` `zs = map (lift_catch P1) xs @ ((Skip,snd(last ((P0, s)#xs)))#ys)` 
      by force    
    then have "(∃xs s' s''. ((Γ,(P, s)#xs) ∈ cptn_mod  ∧ 
             ((zs=(map (lift_catch Q) xs) ∨
             ((fst(((P, s)#xs)!length xs)=Throw ∧
               snd(last ((P, s)#xs)) = Normal s' ∧  s=Normal s''∧
               (∃ys. (Γ,(Q, snd(((P, s)#xs)!length xs))#ys) ∈ cptn_mod ∧ 
                zs=(map (lift_catch Q) xs)@((Q, snd(((P, s)#xs)!length xs))#ys)))) ∨
                (∃ys. ((fst(((P, s)#xs)!length xs)=Skip ∧ (Γ,(Skip,snd(last ((P, s)#xs)))#ys) ∈ cptn_mod ∧                 
                 zs=(map (lift_catch Q) xs)@((Skip,snd(last ((P0, s)#xs)))#ys))))))))"
    using P0cptn  `P0 = P ∧ P1 = Q ∧ s = sa ∧ zs = zsa`  last  CptnModCatch2.hyps(4) by blast 
   } 
   thus ?thesis by auto
  qed
next
  case (CptnModCatch3 Γ P0 s xs s' P1 ys zs)
  from CptnModCatch3.hyps(3)  
  have last:"fst (((P0, Normal s) # xs) ! length xs) = Throw" 
       by (simp add: last_length) 
  from CptnModCatch3.hyps(4) 
  have lastnormal:"snd (last ((P0, Normal s) # xs)) = Normal s'"
      by (simp add: last_length)
  have P0cptn:"(Γ, (P0, Normal s) # xs) ∈ cptn_mod" by fact
  from CptnModCatch3.hyps(5) have P1cptn:"(Γ, (P1, snd (((P0, Normal s) # xs) ! length xs)) # ys) ∈ cptn_mod"
      by (simp add: last_length)          
  then have "zs = map (lift_catch P1) xs @ (P1, snd (last ((P0, Normal s) # xs))) # ys" by (simp add:CptnModCatch3.hyps)
  show ?case
  proof -{
    fix sa P Q zsa    
    assume eq:"(Catch P0 P1, Normal s) # zs = (Catch P Q, Normal sa) # zsa"
    then have "P0 =P ∧ P1 = Q ∧ Normal s= Normal sa ∧ zs=zsa" by auto     
    have "last ((P0, Normal s) # xs) = ((P, Normal sa) # xs) ! length xs"
      by (simp add: `P0 = P ∧ P1 = Q ∧ Normal s = Normal sa ∧ zs = zsa` last_length)
    then have "zsa = map (lift_catch Q) xs @ (Q, snd (((P, Normal sa) # xs) ! length xs)) # ys"
      using `P0 = P ∧ P1 = Q ∧ Normal s = Normal sa ∧ zs = zsa` `zs = map (lift_catch P1) xs @ (P1, snd (last ((P0, Normal s) # xs))) # ys` by force
    then have "(Γ, (P, Normal s) # xs) ∈ cptn_mod ∧ (fst(((P, Normal s)#xs)!length xs)=Throw ∧
               snd(last ((P, Normal s)#xs)) = Normal s' ∧ 
              (∃ys. (Γ,(Q, snd(((P, Normal s)#xs)!length xs))#ys) ∈ cptn_mod ∧ 
              zs=(map (lift_catch Q) xs)@((Q, snd(((P, Normal s)#xs)!length xs))#ys)))" 
      using lastnormal P1cptn P0cptn `P0 = P ∧ P1 = Q ∧ Normal s = Normal sa ∧ zs = zsa` last 
       by auto    
    }note this [of P0 P1 s zs] thus ?thesis by blast qed
next
  case (CptnModEnv Γ P s t xs)  
  then have step:"(Γ, (P, t) # xs) ∈ cptn_mod" by auto  
  have step_e: "Γ⊢c (P, s) →e (P, t)" using CptnModEnv by auto
  show ?case     
    proof (cases P)
      case (Catch P1 P2) 
      then have eq_P_Catch:"(P, t) # xs = (LanguageCon.com.Catch P1 P2, t) # xs" by auto      
      then  obtain xsa t' t'' where 
         p1:"(Γ, (P1, t) # xsa) ∈ cptn_mod" and p2:"
     (xs = map (lift_catch P2) xsa ∨
      fst (((P1, t) # xsa) ! length xsa) = LanguageCon.com.Throw ∧
      snd (last ((P1, t) # xsa)) = Normal t' ∧
      t = Normal t'' ∧
      (∃ys. (Γ, (P2, snd (((P1, t) # xsa) ! length xsa)) # ys) ∈ cptn_mod ∧
            xs =
            map (lift_catch P2) xsa @
            (P2, snd (((P1, t) # xsa) ! length xsa)) # ys) ∨
            fst (((P1, t) # xsa) ! length xsa) = LanguageCon.com.Skip ∧
            (∃ys.(Γ,(Skip,snd(last ((P1, t)#xsa)))#ys) ∈ cptn_mod ∧ 
            xs = map (lift_catch P2) xsa @
            ((LanguageCon.com.Skip, snd (last ((P1, t) # xsa)))#ys)))" 
        using CptnModEnv(3) by auto
      have all_step:"(Γ, (P1, s)#((P1, t) # xsa)) ∈ cptn_mod"
        by (metis p1 Env Env_n cptn_mod.CptnModEnv env_normal_s step_e)       
      show ?thesis using p2 
      proof  
        assume "xs = map (lift_catch P2) xsa"
        have "(P, t) # xs = map (lift_catch P2) ((P1, t) # xsa)"
          by (simp add: `xs = map (lift_catch P2) xsa` lift_catch_def local.Catch)
        thus ?thesis using all_step eq_P_Catch by fastforce
      next 
        assume 
         "fst (((P1, t) # xsa) ! length xsa) = LanguageCon.com.Throw ∧
          snd (last ((P1, t) # xsa)) = Normal t' ∧
          t = Normal t'' ∧
          (∃ys. (Γ, (P2, snd (((P1, t) # xsa) ! length xsa)) # ys) ∈ cptn_mod ∧
                xs =
                map (lift_catch P2) xsa @
                (P2, snd (((P1, t) # xsa) ! length xsa)) # ys) ∨
                fst (((P1, t) # xsa) ! length xsa) = LanguageCon.com.Skip ∧
           (∃ys. (Γ,(Skip,snd(last ((P1, t)#xsa)))#ys) ∈ cptn_mod ∧ 
            xs = map (lift_catch P2) xsa @
            ((LanguageCon.com.Skip, snd (last ((P1, t) # xsa)))#ys))"      
         then show ?thesis 
         proof
           assume 
            a1:"fst (((P1, t) # xsa) ! length xsa) = LanguageCon.com.Throw ∧
             snd (last ((P1, t) # xsa)) = Normal t' ∧
             t = Normal t'' ∧
             (∃ys. (Γ, (P2, snd (((P1, t) # xsa) ! length xsa)) # ys) ∈ cptn_mod ∧
                xs = map (lift_catch P2) xsa @
                       (P2, snd (((P1, t) # xsa) ! length xsa)) # ys)"
            then obtain ys where p2_exec:"(Γ, (P2, snd (((P1, t) # xsa) ! length xsa)) # ys) ∈ cptn_mod ∧
                xs = map (lift_catch P2) xsa @
                       (P2, snd (((P1, t) # xsa) ! length xsa)) # ys" 
            by fastforce
            from a1 obtain t1 where t_normal: "t=Normal t1" 
              using env_normal_s'_normal_s by blast
            have f1:"fst (((P1, s)#(P1, t) # xsa) ! length ((P1, t)#xsa)) = LanguageCon.com.Throw"
              using a1 by fastforce
             from a1 have last_normal: "snd (last ((P1, s)#(P1, t) # xsa)) = Normal t'"
               by fastforce 
             then have p2_long_exec: "(Γ, (P2, snd (((P1, s)#(P1, t) # xsa) ! length ((P1, s)#xsa))) # ys) ∈ cptn_mod ∧
                (P, t)#xs = map (lift_catch P2) ((P1, t) # xsa) @
                       (P2, snd (((P1, s)#(P1, t) # xsa) ! length ((P1, s)#xsa))) # ys" using p2_exec 
                by (simp add: lift_catch_def local.Catch)                  
             thus ?thesis using  a1 f1 last_normal all_step eq_P_Catch 
               by (clarify, metis (no_types) list.size(4) not_step_c_env  step_e)            
           next
           assume 
            as1:"fst (((P1, t) # xsa) ! length xsa) = LanguageCon.com.Skip ∧
           (∃ys. (Γ,(Skip,snd(last ((P1, t)#xsa)))#ys) ∈ cptn_mod ∧ 
            xs = map (lift_catch P2) xsa @
            ((LanguageCon.com.Skip, snd (last ((P1, t) # xsa)))#ys))"               
            then obtain ys where p1:"(Γ,(Skip,snd(last ((P1, t)#xsa)))#ys) ∈ cptn_mod ∧ 
                         (P, t)#xs = map (lift_catch P2) ((P1, t) # xsa) @
                          ((LanguageCon.com.Skip, snd (last ((P1, t) # xsa)))#ys)"
            proof -
              assume a1: "⋀ys. (Γ, (LanguageCon.com.Skip, snd (last ((P1, t) # xsa))) # ys) ∈ cptn_mod ∧ (P, t) # xs = map (lift_catch P2) ((P1, t) # xsa) @ (LanguageCon.com.Skip, snd (last ((P1, t) # xsa))) # ys ⟹ thesis"
              have "(LanguageCon.com.Catch P1 P2, t) # map (lift_catch P2) xsa = map (lift_catch P2) ((P1, t) # xsa)"
                by (simp add: lift_catch_def)
              thus ?thesis
                using a1 as1 eq_P_Catch by moura
            qed            
            from as1 have p2: "fst (((P1, s)#(P1, t) # xsa) ! length ((P1, t) #xsa)) = LanguageCon.com.Skip"
                 by fastforce                              
            thus ?thesis using p1 all_step eq_P_Catch by fastforce
          qed
      qed
    qed (auto)
qed(force+)


definition seq_cond 
where
"seq_cond zs Q xs P s s'' s' Γ ≡ (zs=(map (lift Q) xs) ∨
             ((fst(((P, s)#xs)!length xs)=Skip ∧ 
               (∃ys. (Γ,(Q, snd(((P, s)#xs)!length xs))#ys) ∈ cptn_mod ∧ 
                zs=(map (lift (Q)) xs)@((Q, snd(((P, s)#xs)!length xs))#ys)))) ∨
             ((fst(((P, s)#xs)!length xs)=Throw ∧ 
                 snd(last ((P, s)#xs)) = Normal s' ∧  s=Normal s''∧
                 (∃ys.  (Γ,(Throw,Normal s')#ys) ∈ cptn_mod ∧
                      zs=(map (lift Q) xs)@((Throw,Normal s')#ys)))))
"

lemma div_seq: assumes cptn_m:"(Γ,list) ∈ cptn_mod"
shows "(∀s P Q zs. list=(Seq P Q, s)#zs ⟶
       (∃xs s' s''. 
          (Γ,(P, s)#xs) ∈ cptn_mod  ∧ 
             seq_cond zs Q xs P s s'' s' Γ))  
            "
unfolding seq_cond_def
using cptn_m
proof (induct rule: cptn_mod.induct)
  case (CptnModOne Γ P s)
  thus ?case using cptn_mod.CptnModOne by blast 
next
  case (CptnModSkip  Γ P s t xs) 
  from CptnModSkip.hyps
  have step: "Γ⊢c (P, s) → (Skip, t)" by auto
  from CptnModSkip.hyps
  have noskip: "~(P=Skip)" using stepc_elim_cases(1) by blast  
  have x: "∀c c1 c2. redex c = Seq c1 c2 ⟹ False"
          using redex_not_Seq by blast
  from CptnModSkip.hyps
  have in_cptn_mod: "(Γ, (Skip, t) # xs) ∈ cptn_mod" by auto  
  then show ?case using CptnModSkip.hyps(2) SmallStepCon.redex_not_Seq by blast
next
  case (CptnModThrow  Γ P s t xs)
  from CptnModThrow.hyps
  have step: "Γ⊢c (P, s) → (Throw, t)" by auto 
  moreover from CptnModThrow.hyps
  have in_cptn_mod: "(Γ, (Throw, t) # xs) ∈ cptn_mod" by auto 
  have no_seq: "∀p1 p2. ¬(P=Seq p1 p2)" using CptnModThrow.hyps(2) redex_not_Seq by auto
  ultimately show ?case by auto   
next 
  case (CptnModCondT Γ P0 s ys b P1)
  thus ?case by auto
next
  case (CptnModCondF Γ P0 s ys b P1)
  thus ?case by auto
next
  case (CptnModSeq1 Γ P0 s xs zs P1)
  thus ?case by blast
next 
  case (CptnModSeq2 Γ P0 s xs P1 ys zs) 
  from CptnModSeq2.hyps(3) last_length have last:"fst (((P0, s) # xs) ! length xs) = Skip" 
       by (simp add: last_length) 
  have P0cptn:"(Γ, (P0, s) # xs) ∈ cptn_mod" by fact
  from CptnModSeq2.hyps(4) have P1cptn:"(Γ, (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∈ cptn_mod"
      by (simp add: last_length)          
  then have "zs = map (lift P1) xs @ (P1, snd (last ((P0, s) # xs))) # ys" by (simp add:CptnModSeq2.hyps)
  show ?case
  proof -{
    fix sa P Q zsa    
    assume eq:"(Seq P0 P1, s) # zs = (Seq P Q, sa) # zsa"
    then have "P0 =P ∧ P1 = Q ∧ s=sa ∧ zs=zsa" by auto 
     have "last ((P0, s) # xs) = ((P, sa) # xs) ! length xs"
            by (simp add: `P0 = P ∧ P1 = Q ∧ s = sa ∧ zs = zsa` last_length)
    then have "zsa = map (lift Q) xs @ (Q, snd (((P, sa) # xs) ! length xs)) # ys"
         using `P0 = P ∧ P1 = Q ∧ s = sa ∧ zs = zsa` `zs = map (lift P1) xs @ (P1, snd (last ((P0, s) # xs))) # ys` 
         by force    
    then have "(∃xs s' s''. (Γ, (P, sa) # xs) ∈ cptn_mod ∧
                        (zsa = map (lift Q) xs ∨              
                         fst (((P, sa) # xs) ! length xs) = Skip ∧
                             (∃ys. (Γ, (Q, snd (((P, sa) # xs) ! length xs)) # ys) ∈ cptn_mod ∧
                             zsa = map (lift Q) xs @ (Q, snd (((P, sa) # xs) ! length xs)) # ys) ∨
                        ((fst(((P, sa)#xs)!length xs)=Throw ∧ 
                          snd(last ((P, sa)#xs)) = Normal s' ∧  s=Normal s''∧
                         (∃ys. (Γ,(Throw,Normal s')#ys) ∈ cptn_mod ∧
                               zsa=(map (lift Q) xs)@((Throw,Normal s')#ys))))))
               " 
        using P0cptn P1cptn  `P0 = P ∧ P1 = Q ∧ s = sa ∧ zs = zsa` last 
        by blast                 
   }  
   thus ?case by auto qed     
next
  case (CptnModSeq3 Γ P0 s xs s' ys zs P1) 
  from CptnModSeq3.hyps(3) 
  have last:"fst (((P0, Normal s) # xs) ! length xs) = Throw" 
       by (simp add: last_length) 
  have P0cptn:"(Γ, (P0, Normal s) # xs) ∈ cptn_mod" by fact
  from CptnModSeq3.hyps(4) 
  have lastnormal:"snd (last ((P0, Normal s) # xs)) = Normal s'"
      by (simp add: last_length)          
  then have "zs = map (lift P1) xs @ ((Throw, Normal s')#ys)" by (simp add:CptnModSeq3.hyps)
  show ?case
  proof -{
    fix sa P Q zsa    
    assume eq:"(Seq P0 P1, Normal s) # zs = (Seq P Q, Normal sa) # zsa"
    then have "P0 =P ∧ P1 = Q ∧ Normal s=Normal sa ∧ zs=zsa" by auto
    then have "(P0, Normal s) = (P, Normal sa)" by auto 
    have "last ((P0, Normal s) # xs) = ((P, Normal sa) # xs) ! length xs"
                    by (simp add: `P0 = P ∧ P1 = Q ∧ Normal s = Normal sa ∧ zs = zsa` last_length)
    then have zsa:"zsa = (map (lift Q) xs)@((Throw,Normal s')#ys)"
                    using `P0 = P ∧ P1 = Q ∧ Normal s = Normal sa ∧ zs = zsa` `zs = map (lift P1) xs @ ((Throw, Normal s')#ys)` 
    by force
    then have a1:"(Γ,(Throw,Normal s')#ys) ∈ cptn_mod" using CptnModSeq3.hyps(5) by blast               
     have "(P, Normal sa::('b, 'c) xstate) = (P0, Normal s)"
    using `P0 = P ∧ P1 = Q ∧ Normal s = Normal sa ∧ zs = zsa` by auto  
    then have "(∃xs s'. (Γ, (P, Normal sa) # xs) ∈ cptn_mod ∧
                        (zsa = map (lift Q) xs ∨              
                         fst (((P,Normal sa) # xs) ! length xs) = Skip ∧
                             (∃ys. (Γ, (Q, snd (((P, Normal sa) # xs) ! length xs)) # ys) ∈ cptn_mod ∧
                             zsa = map (lift Q) xs @ (Q, snd (((P, Normal sa) # xs) ! length xs)) # ys) ∨
                        ((fst(((P, Normal sa)#xs)!length xs)=Throw ∧ 
                          snd(last ((P, Normal sa)#xs)) = Normal s' ∧
                          (∃ys. (Γ,(Throw,Normal s')#ys) ∈ cptn_mod ∧
                          zsa=(map (lift Q) xs)@((Throw,Normal s')#ys))))))"
     using P0cptn zsa a1 last lastnormal 
       by blast 
   }  
   thus ?thesis by auto qed  
next 
  case (CptnModEnv Γ P s t zs) 
  then have step:"(Γ, (P, t) # zs) ∈ cptn_mod" by auto
  have step_e: "Γ⊢c (P, s) →e (P, t)" using CptnModEnv by auto  
  show ?case     
    proof (cases P)
      case (Seq P1 P2) 
      then have eq_P:"(P, t) # zs = (LanguageCon.com.Seq P1 P2, t) # zs" by auto      
      then  obtain xs t' t'' where 
         p1:"(Γ, (P1, t) # xs) ∈ cptn_mod" and p2:"
     (zs = map (lift P2) xs ∨
      fst (((P1, t) # xs) ! length xs) = LanguageCon.com.Skip ∧      
      (∃ys. (Γ, (P2, snd (((P1, t) # xs) ! length xs)) # ys) ∈ cptn_mod ∧
            zs =
            map (lift P2) xs @
            (P2, snd (((P1, t) # xs) ! length xs)) # ys) ∨
      fst (((P1, t) # xs) ! length xs) = LanguageCon.com.Throw ∧
      snd (last ((P1, t) # xs)) = Normal t' ∧
      t = Normal t'' ∧ (∃ys. (Γ,(Throw,Normal t')#ys) ∈ cptn_mod ∧ 
      zs =
      map (lift P2) xs @
      ((LanguageCon.com.Throw, Normal t')#ys))) " 
        using CptnModEnv(3) by auto
      have all_step:"(Γ, (P1, s)#((P1, t) # xs)) ∈ cptn_mod"
       by (metis p1 Env Env_n cptn_mod.CptnModEnv env_normal_s step_e)         
      show ?thesis using p2 
      proof  
        assume "zs = map (lift P2) xs"
        have "(P, t) # zs = map (lift P2) ((P1, t) # xs)"
          by (simp add: `zs = map (lift P2) xs` lift_def local.Seq)
        thus ?thesis using all_step eq_P by fastforce
      next 
        assume 
         "fst (((P1, t) # xs) ! length xs) = LanguageCon.com.Skip ∧      
         (∃ys. (Γ, (P2, snd (((P1, t) # xs) ! length xs)) # ys) ∈ cptn_mod ∧
            zs = map (lift P2) xs @ (P2, snd (((P1, t) # xs) ! length xs)) # ys) ∨
          fst (((P1, t) # xs) ! length xs) = LanguageCon.com.Throw ∧
           snd (last ((P1, t) # xs)) = Normal t' ∧
           t = Normal t'' ∧ (∃ys. (Γ,(Throw,Normal t')#ys) ∈ cptn_mod ∧ 
           zs = map (lift P2) xs @ ((LanguageCon.com.Throw, Normal t')#ys))"
         then show ?thesis 
         proof
           assume 
            a1:"fst (((P1, t) # xs) ! length xs) = LanguageCon.com.Skip ∧      
               (∃ys. (Γ, (P2, snd (((P1, t) # xs) ! length xs)) # ys) ∈ cptn_mod ∧
               zs = map (lift P2) xs @ (P2, snd (((P1, t) # xs) ! length xs)) # ys)"
               from a1 obtain ys where 
                   p2_exec:"(Γ, (P2, snd (((P1, t) # xs) ! length xs)) # ys) ∈ cptn_mod ∧
                       zs = map (lift P2) xs @
                       (P2, snd (((P1, t) # xs) ! length xs)) # ys" 
                by auto 
               have f1:"fst (((P1, s)#(P1, t) # xs) ! length ((P1, t)#xs)) = LanguageCon.com.Skip"
                 using a1 by fastforce             
               then have p2_long_exec: 
                 "(Γ, (P2, snd (((P1, s)#(P1, t) # xs) ! length ((P1, t)#xs))) # ys) ∈ cptn_mod ∧
                  (P, t)#zs = map (lift P2) ((P1, t) # xs) @
                       (P2, snd (((P1, s)#(P1, t) # xs) ! length ((P1, t)#xs))) # ys" 
             using p2_exec by (simp add: lift_def local.Seq)     
             thus ?thesis using a1 f1 all_step eq_P by blast            
           next
           assume 
            a1:"fst (((P1, t) # xs) ! length xs) = LanguageCon.com.Throw ∧
            snd (last ((P1, t) # xs)) = Normal t' ∧ t = Normal t'' ∧ 
          (∃ys. (Γ,(Throw,Normal t')#ys) ∈ cptn_mod ∧ 
           zs = map (lift P2) xs @ ((LanguageCon.com.Throw, Normal t')#ys))"             
             then have last_throw:
               "fst (((P1, s)#(P1, t) # xs) ! length ((P1, t) #xs)) = LanguageCon.com.Throw" 
               by fastforce
             from a1 have last_normal: "snd (last ((P1, s)#(P1, t) # xs)) = Normal t'"
               by fastforce
             have seq_lift:
               "(LanguageCon.com.Seq P1 P2, t) # map (lift P2) xs = map (lift P2) ((P1, t) # xs)"
                by (simp add: a1 lift_def)             
            thus  ?thesis using a1 last_throw last_normal all_step eq_P         
              by (clarify, metis (no_types, lifting) append_Cons env_normal_s'_normal_s  step_e)                 
          qed
      qed
    qed (auto) 
qed (force)+


lemma cptn_onlyif_cptn_mod_aux:
assumes stepseq:"Γ⊢c (P, s) → (Q,t)" and
        stepmod:"(Γ,(Q,t)#xs) ∈ cptn_mod"
shows "(Γ,(P,s)#(Q,t)#xs) ∈ cptn_mod"
using stepseq stepmod
proof (induct arbitrary: xs)
  case (Basicc f s) 
  thus ?case by (simp add: cptn_mod.CptnModSkip stepc.Basicc)
next
  case (Specc s t r)
  thus ?case by (simp add: cptn_mod.CptnModSkip stepc.Specc) 
next
  case (SpecStuckc s r)
  thus ?case by (simp add: cptn_mod.CptnModSkip stepc.SpecStuckc)
next
  case (Guardc s g f c)
  thus ?case by (simp add: cptn_mod.CptnModGuard) 
next
  case (GuardFaultc)
  thus ?case by (simp add: cptn_mod.CptnModSkip stepc.GuardFaultc)
next
  case (Seqc c1 s c1' s' c2)
  have step: "Γ⊢c (c1, s) → (c1', s')" by (simp add: Seqc.hyps(1))
  then have nsc1: "c1≠Skip" using stepc_elim_cases(1) by blast 
  have assum: "(Γ, (Seq c1' c2, s') # xs) ∈ cptn_mod" using Seqc.prems by blast
  have divseq:"(∀s P Q zs. (Seq c1' c2, s') # xs=(Seq P Q, s)#zs ⟶
                (∃xs sv' sv''. ((Γ,(P, s)#xs) ∈ cptn_mod  ∧ 
                           (zs=(map (lift Q) xs) ∨
                           ((fst(((P, s)#xs)!length xs)=Skip ∧ 
                             (∃ys. (Γ,(Q, snd(((P, s)#xs)!length xs))#ys) ∈ cptn_mod ∧ 
                              zs=(map (lift (Q)) xs)@((Q, snd(((P, s)#xs)!length xs))#ys)))) ∨
                           ((fst(((P, s)#xs)!length xs)=Throw ∧ 
                               snd(last ((P, s)#xs)) = Normal sv' ∧  s'=Normal sv''∧
                             (∃ys.  (Γ,(Throw,Normal sv')#ys) ∈ cptn_mod ∧
                              zs=(map (lift Q) xs)@((Throw,Normal sv')#ys))
                               ))))
                            
                 ))"  using div_seq [OF assum] unfolding seq_cond_def by auto
   {fix sa P Q zsa
       assume ass:"(Seq c1' c2, s') # xs = (Seq P Q, sa) # zsa"
       then have eqs:"c1' = P ∧ c2 = Q ∧ s' = sa ∧ xs = zsa" by auto
       then have "(∃xs sv' sv''. (Γ, (P, sa) # xs) ∈ cptn_mod ∧
                        (zsa = map (lift Q) xs ∨              
                         fst (((P, sa) # xs) ! length xs) = Skip ∧
                             (∃ys. (Γ, (Q, snd (((P, sa) # xs) ! length xs)) # ys) ∈ cptn_mod ∧
                             zsa = map (lift Q) xs @ (Q, snd (((P, sa) # xs) ! length xs)) # ys) ∨
                        ((fst(((P, sa)#xs)!length xs)=Throw ∧ 
                          snd(last ((P, sa)#xs)) = Normal sv' ∧  s'=Normal sv''∧
                          (∃ys.  (Γ,(Throw,Normal sv')#ys) ∈ cptn_mod ∧
                              zsa=(map (lift Q) xs)@((Throw,Normal sv')#ys))))))" 
             using ass divseq by blast
    } note conc=this [of c1' c2 s' xs]    
     then obtain xs' sa' sa''
       where  split:"(Γ, (c1', s') # xs') ∈ cptn_mod ∧
                     (xs = map (lift c2) xs' ∨                   
                     fst (((c1', s') # xs') ! length xs') = Skip ∧
                        (∃ys. (Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod ∧
                         xs = map (lift c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∨
                     ((fst(((c1', s')#xs')!length xs')=Throw ∧ 
                         snd(last ((c1', s')#xs')) = Normal sa' ∧ s'=Normal sa''∧
                         (∃ys.  (Γ,(Throw,Normal sa')#ys) ∈ cptn_mod ∧
                              xs=(map (lift c2) xs')@((Throw,Normal sa')#ys))
                         )))"  by blast
  then have "(xs = map (lift c2) xs' ∨                   
                     fst (((c1', s') # xs') ! length xs') = Skip ∧
                        (∃ys. (Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod ∧
                         xs = map (lift c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∨
                     ((fst(((c1', s')#xs')!length xs')=Throw ∧ 
                         snd(last ((c1', s')#xs')) = Normal sa' ∧ s'=Normal sa''∧
                         (∃ys.  (Γ,(Throw,Normal sa')#ys) ∈ cptn_mod ∧
                              xs=(map (lift c2) xs')@((Throw,Normal sa')#ys)))))" by auto
  thus ?case 
  proof{           
       assume c1'nonf:"xs = map (lift c2) xs'"
       then have c1'cptn:"(Γ, (c1', s') # xs') ∈ cptn_mod" using split by blast
       then have induct_step: "(Γ, (c1, s) # (c1', s')#xs') ∈ cptn_mod"
         using Seqc.hyps(2)  by blast
       then have "(Seq c1' c2, s')#xs = map (lift c2) ((c1', s')#xs')"
            using c1'nonf
            by (simp add: CptnModSeq1 lift_def)
       thus ?thesis 
            using c1'nonf c1'cptn induct_step by (auto simp add: CptnModSeq1)
    next
      assume "fst (((c1', s') # xs') ! length xs') = Skip ∧
              (∃ys. (Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod ∧
                  xs = map (lift c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∨
             ((fst(((c1', s')#xs')!length xs')=Throw ∧ 
                snd(last ((c1', s')#xs')) = Normal sa' ∧  s'=Normal sa''∧
                (∃ys.  (Γ,(Throw,Normal sa')#ys) ∈ cptn_mod ∧
                              xs=(map (lift c2) xs')@((Throw,Normal sa')#ys))))"  
      thus ?thesis
      proof
         assume assth:"fst (((c1', s') # xs') ! length xs') = Skip ∧
              (∃ys. (Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod ∧
                  xs = map (lift c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys)"
         then obtain ys 
             where split':"(Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod ∧
                  xs = map (lift c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys"
             by auto    
         then have c1'cptn:"(Γ, (c1', s') # xs') ∈ cptn_mod" using split by blast
         then have induct_step: "(Γ, (c1, s) # (c1', s')#xs') ∈ cptn_mod"
         using Seqc.hyps(2)  by blast                
         then have seqmap:"(Seq c1 c2, s)#(Seq c1' c2, s')#xs = map (lift c2) ((c1,s)#(c1', s')#xs') @ (c2, snd (((c1', s') # xs') ! length xs')) # ys"
        using split'   
         by (simp add: CptnModSeq2 lift_def) 
        then have lastc1:"last ((c1, s) # (c1', s') # xs') = ((c1', s') # xs') ! length xs'"
          by (simp add: last_length) 
        then have lastc1skip:"fst (last ((c1, s) # (c1', s') # xs')) = Skip" 
          using assth by fastforce          
        thus ?thesis 
           using seqmap split' last_length cptn_mod.CptnModSeq2 
                 induct_step lastc1 lastc1skip 
           by fastforce
    next
        assume assm:"((fst(((c1', s')#xs')!length xs')=Throw ∧ 
                snd(last ((c1', s')#xs')) = Normal sa' ∧  s'=Normal sa''∧
                (∃ys.  (Γ,(Throw,Normal sa')#ys) ∈ cptn_mod ∧
                 xs=(map (lift c2) xs')@((Throw,Normal sa')#ys))))"
        then have s'eqsa'': "s'=Normal sa''" by auto
        then have snormal: "∃ns. s=Normal ns" by (metis Seqc.hyps(1) step_Abrupt_prop step_Fault_prop step_Stuck_prop xstate.exhaust)
        then have c1'cptn:"(Γ, (c1', s') # xs') ∈ cptn_mod" using split by blast        
        then have induct_step: "(Γ, (c1, s) # (c1', s')#xs') ∈ cptn_mod"
        using Seqc.hyps(2)  by blast 
        then obtain ys where seqmap:"(Seq c1' c2, s')#xs = (map (lift c2) ((c1', s')#xs'))@((Throw,Normal sa')#ys)"
        using assm
        proof -
          assume a1: "⋀ys. (LanguageCon.com.Seq c1' c2, s') # xs = map (lift c2) ((c1', s') # xs') @ (LanguageCon.com.Throw, Normal sa') # ys ⟹ thesis"
          have "(LanguageCon.com.Seq c1' c2, Normal sa'') # map (lift c2) xs' = map (lift c2) ((c1', s') # xs')"
            by (simp add: assm lift_def)
          thus ?thesis
            using a1 assm by moura
        qed  
        then have lastc1:"last ((c1, s) # (c1', s') # xs') = ((c1', s') # xs') ! length xs'"
                   by (simp add: last_length) 
        then have lastc1skip:"fst (last ((c1, s) # (c1', s') # xs')) = Throw" 
             using assm by fastforce
        then have "snd (last ((c1, s) # (c1', s') # xs')) = Normal sa'"
             using assm by force        
        thus ?thesis
            using assm c1'cptn induct_step lastc1skip snormal seqmap s'eqsa'' 
            by (auto simp add:cptn_mod.CptnModSeq3)
   qed
  }qed    
          
next
  case (SeqSkipc c2 s xs)
  have c2incptn:"(Γ, (c2, s) # xs) ∈ cptn_mod" by fact
  then have 1:"(Γ, [(Skip, s)]) ∈ cptn_mod" by (simp add: cptn_mod.CptnModOne)
  then have 2:"fst(last ([(Skip, s)])) = Skip" by fastforce
  then have 3:"(Γ,(c2, snd(last [(Skip, s)]))#xs) ∈ cptn_mod" 
      using c2incptn by auto  
  then have "(c2,s)#xs=(map (lift c2) [])@(c2, snd(last [(Skip, s)]))#xs" 
       by (auto simp add:lift_def)
  thus ?case using 1 2 3 by (simp add: CptnModSeq2)
next
  case (SeqThrowc c2 s xs)  
  have "(Γ, [(Throw, Normal s)]) ∈ cptn_mod" by (simp add: cptn_mod.CptnModOne) 
  then obtain ys where ys_nil:"ys=[]" and last:"(Γ, (Throw, Normal s)#ys)∈ cptn_mod"
   by auto
  moreover have "fst (last ((Throw, Normal s)#ys)) = Throw" using ys_nil last by auto
  moreover have "snd (last ((Throw, Normal s)#ys)) = Normal s" using ys_nil last by auto
  moreover from ys_nil have "(map (lift c2) ys) = []" by auto  
  ultimately show ?case using SeqThrowc.prems cptn_mod.CptnModSeq3 by fastforce  
next
  case (CondTruec s b c1 c2)
  thus ?case by (simp add: cptn_mod.CptnModCondT)
next
  case (CondFalsec s b c1 c2)
  thus ?case by (simp add: cptn_mod.CptnModCondF)
next
 case (WhileTruec s1 b c)
 have sinb: "s1∈b" by fact
 have SeqcWhile: "(Γ, (Seq c (While b c), Normal s1) # xs) ∈ cptn_mod" by fact  
 have divseq:"(∀s P Q zs. (Seq c (While b c), Normal s1) # xs=(Seq P Q, s)#zs ⟶
                (∃xs s'. ((Γ,(P, s)#xs) ∈ cptn_mod  ∧ 
                           (zs=(map (lift Q) xs) ∨
                           ((fst(((P, s)#xs)!length xs)=Skip ∧ 
                             (∃ys. (Γ,(Q, snd(((P, s)#xs)!length xs))#ys) ∈ cptn_mod ∧ 
                              zs=(map (lift (Q)) xs)@((Q, snd(((P, s)#xs)!length xs))#ys)))) ∨
                           ((fst(((P, s)#xs)!length xs)=Throw ∧ 
                               snd(last ((P, s)#xs)) = Normal s' ∧
                              (∃ys.  (Γ,(Throw,Normal s')#ys) ∈ cptn_mod ∧
                      zs=(map (lift Q) xs)@((Throw,Normal s')#ys))))))                            
                 ))" using div_seq [OF SeqcWhile] by (auto simp add: seq_cond_def)
{fix sa P Q zsa
       assume ass:"(Seq c (While b c), Normal s1) # xs = (Seq P Q, sa) # zsa"
       then have eqs:"c = P ∧ (While b c) = Q ∧ Normal s1 = sa ∧ xs = zsa" by auto
       then have "(∃xs s'. (Γ, (P, sa) # xs) ∈ cptn_mod ∧
                        (zsa = map (lift Q) xs ∨              
                         fst (((P, sa) # xs) ! length xs) = Skip ∧
                             (∃ys. (Γ, (Q, snd (((P, sa) # xs) ! length xs)) # ys) ∈ cptn_mod ∧
                             zsa = map (lift Q) xs @ (Q, snd (((P, sa) # xs) ! length xs)) # ys) ∨
                        ((fst(((P, sa)#xs)!length xs)=Throw ∧ 
                          snd(last ((P, sa)#xs)) = Normal s' ∧
                          (∃ys.  (Γ,(Throw,Normal s')#ys) ∈ cptn_mod ∧
                      zsa=(map (lift Q) xs)@((Throw,Normal s')#ys))
                        ))))" 
             using ass divseq by auto
    } note conc=this [of c "While b c" "Normal s1" xs] 
   then obtain xs' s' 
        where  split:"(Γ, (c, Normal s1) # xs') ∈ cptn_mod ∧
     (xs = map (lift (While b c)) xs' ∨
      fst (((c, Normal s1) # xs') ! length xs') = Skip ∧
      (∃ys. (Γ, (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys)
            ∈ cptn_mod ∧
            xs =
            map (lift (While b c)) xs' @
            (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys) ∨
      fst (((c, Normal s1) # xs') ! length xs') = Throw ∧
      snd (last ((c, Normal s1) # xs')) = Normal s' ∧ 
      (∃ys.  (Γ, ((Throw, Normal s')#ys)) ∈ cptn_mod ∧
      xs = map (lift (While b c)) xs' @ ((Throw, Normal s')#ys)))"  by auto
 then have "(xs = map (lift (While b c)) xs' ∨
            fst (((c, Normal s1) # xs') ! length xs') = Skip ∧
            (∃ys. (Γ, (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys)
                  ∈ cptn_mod ∧
                  xs =
                  map (lift (While b c)) xs' @
                  (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys) ∨
            fst (((c, Normal s1) # xs') ! length xs') = Throw ∧
            snd (last ((c, Normal s1) # xs')) = Normal s' ∧
            (∃ys.  (Γ, ((Throw, Normal s')#ys)) ∈ cptn_mod ∧
          xs = map (lift (While b c)) xs' @ ((Throw, Normal s')#ys)))" ..
 thus ?case
 proof{ 
   assume 1:"xs = map (lift (While b c)) xs'"  
   have 3:"(Γ, (c, Normal s1) # xs') ∈ cptn_mod" using split by auto   
   then show ?thesis using "1" cptn_mod.CptnModWhile1 sinb by fastforce 
 next
   assume "fst (((c, Normal s1) # xs') ! length xs') = Skip ∧
          (∃ys. (Γ, (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys)
                ∈ cptn_mod ∧
                xs =
                map (lift (While b c)) xs' @
                (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys) ∨
          fst (((c, Normal s1) # xs') ! length xs') = Throw ∧
          snd (last ((c, Normal s1) # xs')) = Normal s' ∧
          (∃ys.  (Γ, ((Throw, Normal s')#ys)) ∈ cptn_mod ∧
          xs = map (lift (While b c)) xs' @ ((Throw, Normal s')#ys))"
   thus ?case
   proof
     assume asm:"fst (((c, Normal s1) # xs') ! length xs') = Skip ∧
             (∃ys. (Γ, (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys)
             ∈ cptn_mod ∧
             xs =
             map (lift (While b c)) xs' @
             (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys)"
     then obtain ys 
       where asm':"(Γ, (While b c, snd (last ((c, Normal s1) # xs'))) # ys)
                   ∈ cptn_mod 
                   ∧ xs = map (lift (While b c)) xs' @
                       (While b c, snd (last ((c, Normal s1) # xs'))) # ys" 
              by (auto simp add: last_length)
     moreover have 3:"(Γ, (c, Normal s1) # xs') ∈ cptn_mod" using split by auto
     moreover from asm have "fst (last ((c, Normal s1) # xs'))  = Skip"
          by (simp add: last_length) 
     ultimately show ?case using sinb by (auto simp add:CptnModWhile2)
   next
    assume asm:" fst (((c, Normal s1) # xs') ! length xs') = Throw ∧
          snd (last ((c, Normal s1) # xs')) = Normal s' ∧
          (∃ys.  (Γ, ((Throw, Normal s')#ys)) ∈ cptn_mod ∧
          xs = map (lift (While b c)) xs' @ ((Throw, Normal s')#ys))"   
     moreover have 3:"(Γ, (c, Normal s1) # xs') ∈ cptn_mod" using split by auto
     moreover from asm have "fst (last ((c, Normal s1) # xs'))  = Throw"
          by (simp add: last_length) 
     ultimately show ?case using sinb by (auto simp add:CptnModWhile3)
   qed
 }qed  
next
 case (WhileFalsec s b c)
 thus ?case by (simp add: cptn_mod.CptnModSkip stepc.WhileFalsec)
next
  case (Awaitc s b c t)
  thus ?case by (simp add: cptn_mod.CptnModSkip stepc.Awaitc)
next
  case (AwaitAbruptc s b c t t')
  thus ?case by (simp add: cptn_mod.CptnModThrow stepc.AwaitAbruptc) 
next
  case (Callc p bdy s)
  thus ?case by (simp add: cptn_mod.CptnModCall) 
next
  case (CallUndefinedc p s)
  thus ?case by (simp add: cptn_mod.CptnModSkip stepc.CallUndefinedc)
next
  case (DynComc c s)
  thus ?case by (simp add: cptn_mod.CptnModDynCom) 
next
  case (Catchc c1 s c1' s' c2)
   have step: "Γ⊢c (c1, s) → (c1', s')" by (simp add: Catchc.hyps(1))
  then have nsc1: "c1≠Skip" using stepc_elim_cases(1) by blast 
  have assum: "(Γ, (Catch c1' c2, s') # xs) ∈ cptn_mod" 
  using Catchc.prems by blast
  have divcatch:"(∀s P Q zs. (Catch c1' c2, s') # xs=(Catch P Q, s)#zs ⟶
  (∃xs s' s''. ((Γ,(P, s)#xs) ∈ cptn_mod  ∧ 
             (zs=(map (lift_catch Q) xs) ∨
             ((fst(((P, s)#xs)!length xs)=Throw ∧
               snd(last ((P, s)#xs)) = Normal s' ∧  s=Normal s''∧
               (∃ys. (Γ,(Q, snd(((P, s)#xs)!length xs))#ys) ∈ cptn_mod ∧ 
                zs=(map (lift_catch Q) xs)@((Q, snd(((P, s)#xs)!length xs))#ys)))) ∨
                ((fst(((P, s)#xs)!length xs)=Skip ∧  
                  (∃ys. (Γ,(Skip,snd(last ((P, s)#xs)))#ys) ∈ cptn_mod ∧                   
                 zs=(map (lift_catch Q) xs)@((Skip,snd(last ((P, s)#xs)))#ys))                
                 ))))
   ))" using div_catch [OF assum] by (auto simp add: catch_cond_def)
   {fix sa P Q zsa
       assume ass:"(Catch c1' c2, s') # xs = (Catch P Q, sa) # zsa"
       then have eqs:"c1' = P ∧ c2 = Q ∧ s' = sa ∧ xs = zsa" by auto
       then have "(∃xs sv' sv''. ((Γ,(P, sa)#xs) ∈ cptn_mod  ∧ 
             (zsa=(map (lift_catch Q) xs) ∨
             ((fst(((P, sa)#xs)!length xs)=Throw ∧
               snd(last ((P, sa)#xs)) = Normal sv' ∧  s'=Normal sv''∧
               (∃ys. (Γ,(Q, snd(((P, sa)#xs)!length xs))#ys) ∈ cptn_mod ∧ 
                zsa=(map (lift_catch Q) xs)@((Q, snd(((P, sa)#xs)!length xs))#ys)))) ∨
                ((fst(((P, sa)#xs)!length xs)=Skip ∧                  
                 (∃ys. (Γ,(Skip,snd(last ((P, sa)#xs)))#ys) ∈ cptn_mod ∧                   
                 zsa=(map (lift_catch Q) xs)@((Skip,snd(last ((P, sa)#xs)))#ys))))))
   )"   using ass divcatch by blast
    } note conc=this [of c1' c2 s' xs]    
     then obtain xs' sa' sa''
       where split:
         "(Γ, (c1', s') # xs') ∈ cptn_mod ∧ 
          (xs = map (lift_catch c2) xs' ∨
          fst (((c1', s') # xs') ! length xs') = Throw ∧
          snd (last ((c1', s') # xs')) = Normal sa' ∧ s' = Normal sa'' ∧
          (∃ys. (Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod ∧
                xs = map (lift_catch c2) xs' @
                (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∨
          fst (((c1', s') # xs') ! length xs') = Skip ∧
          (∃ys. (Γ,(Skip,snd(last ((c1', s')#xs')))#ys) ∈ cptn_mod ∧                   
                 xs=(map (lift_catch c2) xs')@((Skip,snd(last ((c1', s')#xs')))#ys)))"          
        by blast
  then have "(xs = map (lift_catch c2) xs' ∨
          fst (((c1', s') # xs') ! length xs') = Throw ∧
          snd (last ((c1', s') # xs')) = Normal sa' ∧ s' = Normal sa'' ∧
          (∃ys. (Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod ∧
                xs = map (lift_catch c2) xs' @
                (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∨
          fst (((c1', s') # xs') ! length xs') = Skip ∧
          (∃ys. (Γ,(Skip,snd(last ((c1', s')#xs')))#ys) ∈ cptn_mod ∧                   
                 xs=(map (lift_catch c2) xs')@((Skip,snd(last ((c1', s')#xs')))#ys)))"          
        by auto
  thus ?case 
  proof{           
       assume c1'nonf:"xs = map (lift_catch c2) xs'"
       then have c1'cptn:"(Γ, (c1', s') # xs') ∈ cptn_mod" using split by blast
       then have induct_step: "(Γ, (c1, s) # (c1', s')#xs') ∈ cptn_mod"
         using Catchc.hyps(2)  by blast
       then have "(Catch c1' c2, s')#xs = map (lift_catch c2) ((c1', s')#xs')"
            using c1'nonf
            by (simp add: CptnModCatch1 lift_catch_def)
       thus ?thesis 
            using c1'nonf c1'cptn induct_step by (auto simp add: CptnModCatch1)
    next
      assume "fst (((c1', s') # xs') ! length xs') = Throw ∧
                snd (last ((c1', s') # xs')) = Normal sa' ∧ s' = Normal sa'' ∧
                (∃ys. (Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod ∧
                xs =map (lift_catch c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∨
               fst (((c1', s') # xs') ! length xs') = Skip ∧
                (∃ys. (Γ,(Skip,snd(last ((c1', s')#xs')))#ys) ∈ cptn_mod ∧                   
                 xs=(map (lift_catch c2) xs')@((Skip,snd(last ((c1', s')#xs')))#ys))"  
      thus ?thesis
      proof
         assume assth:
               "fst (((c1', s') # xs') ! length xs') = Throw ∧
                snd (last ((c1', s') # xs')) = Normal sa' ∧ s' = Normal sa'' ∧
                (∃ys. (Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod ∧
                xs =map (lift_catch c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys)"
             then have s'eqsa'': "s'=Normal sa''" by auto
             then have snormal: "∃ns. s=Normal ns" by (metis Catchc.hyps(1) step_Abrupt_prop step_Fault_prop step_Stuck_prop xstate.exhaust)
             then obtain ys 
             where split':"(Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod ∧
                xs =map (lift_catch c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys"
                using assth by auto    
         then have c1'cptn:"(Γ, (c1', s') # xs') ∈ cptn_mod" 
              using split by blast
         then have induct_step: "(Γ, (c1, s) # (c1', s')#xs') ∈ cptn_mod"
              using Catchc.hyps(2)  by blast                
         then have seqmap:"(Catch c1 c2, s)#(Catch c1' c2, s')#xs = map (lift_catch c2) ((c1,s)#(c1', s')#xs') @ (c2, snd (((c1', s') # xs') ! length xs')) # ys"
              using split' by (simp add: CptnModCatch3 lift_catch_def) 
        then have lastc1:"last ((c1, s) # (c1', s') # xs') = ((c1', s') # xs') ! length xs'"
             by (simp add: last_length) 
        then have lastc1skip:"fst (last ((c1, s) # (c1', s') # xs')) = Throw" 
             using assth by fastforce
        then have "snd (last ((c1, s) # (c1', s') # xs')) = Normal sa'"
             using assth by force
        thus ?thesis using snormal seqmap s'eqsa'' split' last_length  cptn_mod.CptnModCatch3 induct_step lastc1 lastc1skip
             by fastforce 
    next
        assume assm:" fst (((c1', s') # xs') ! length xs') = Skip ∧ 
                       (∃ys. (Γ,(Skip,snd(last ((c1', s')#xs')))#ys) ∈ cptn_mod ∧                   
                      xs=(map (lift_catch c2) xs')@((Skip,snd(last ((c1', s')#xs')))#ys))"
        then have c1'cptn:"(Γ, (c1', s') # xs') ∈ cptn_mod" using split by blast
        then have induct_step: "(Γ, (c1, s) # (c1', s')#xs') ∈ cptn_mod"
        using Catchc.hyps(2)  by blast 
        then have "map (lift_catch c2) ((c1', s') # xs') = (Catch c1' c2, s') # map (lift_catch c2) xs'" 
          by (auto simp add: lift_catch_def)
        then obtain ys 
             where seqmap:"(Catch c1' c2, s')#xs = (map (lift_catch c2) ((c1', s')#xs'))@((Skip,snd(last ((c1', s')#xs')))#ys)"
        using assm by fastforce
        then have lastc1:"last ((c1, s) # (c1', s') # xs') = ((c1', s') # xs') ! length xs'"
                   by (simp add: last_length) 
        then have lastc1skip:"fst (last ((c1, s) # (c1', s') # xs')) = Skip" 
             using assm by fastforce
        then have "snd (last ((c1, s) # (c1', s') # xs')) = snd (last ((c1', s') # xs'))"
             using assm by force
        thus ?thesis 
            using assm c1'cptn induct_step lastc1skip seqmap by (auto simp add:cptn_mod.CptnModCatch2)
    qed
  }qed              
next
  case (CatchThrowc c2 s)
  have c2incptn:"(Γ, (c2, Normal s) # xs) ∈ cptn_mod" by fact
  then have 1:"(Γ, [(Throw, Normal s)]) ∈ cptn_mod" by (simp add: cptn_mod.CptnModOne)
  then have 2:"fst(last ([(Throw, Normal s)])) = Throw" by fastforce
  then have 3:"(Γ,(c2, snd(last [(Throw, Normal s)]))#xs) ∈ cptn_mod" 
      using c2incptn by auto  
  then have "(c2,Normal s)#xs=(map (lift c2) [])@(c2, snd(last [(Throw, Normal s)]))#xs" 
       by (auto simp add:lift_def)
  thus ?case using 1 2 3 by (simp add: CptnModCatch3)
next
  case (CatchSkipc c2 s)
  have "(Γ, [(Skip, s)]) ∈ cptn_mod" by (simp add: cptn_mod.CptnModOne)
  then obtain ys where ys_nil:"ys=[]" and last:"(Γ, (Skip,  s)#ys)∈ cptn_mod"
    by auto
  moreover have "fst (last ((Skip,  s)#ys)) = Skip" using ys_nil last by auto
  moreover have "snd (last ((Skip,  s)#ys)) = s" using ys_nil last by auto
  moreover from ys_nil have "(map (lift_catch c2) ys) = []" by auto  
  ultimately show ?case using CatchSkipc.prems by simp (simp add: cptn_mod.CptnModCatch2 ys_nil)
next
  case (FaultPropc c f)
  thus ?case by (simp add: cptn_mod.CptnModSkip stepc.FaultPropc) 
next
  case (AbruptPropc c f)
  thus ?case by (simp add: cptn_mod.CptnModSkip stepc.AbruptPropc)
next
  case (StuckPropc c)
  thus ?case by (simp add: cptn_mod.CptnModSkip stepc.StuckPropc)
qed

lemma cptn_onlyif_cptn_mod: 
assumes cptn_asm:"(Γ,c) ∈ cptn"
shows  "(Γ,c) ∈ cptn_mod"
using cptn_asm
proof (induct) 
 case CptnOne thus ?case by (rule CptnModOne)
next
 case (CptnEnv Γ P t xs s) thus ?case by (simp add: cptn_mod.CptnModEnv) 
next
 case CptnComp thus ?case
 by (simp add: cptn_onlyif_cptn_mod_aux)
qed

lemma lift_is_cptn: 
assumes cptn_asm:"(Γ,c)∈cptn"
shows "(Γ,map (lift P) c) ∈ cptn"
using cptn_asm
proof (induct)
 case CptnOne thus ?case using cptn.simps by fastforce
next
  case (CptnEnv Γ P s t xs) thus ?case 
      by (cases rule:step_e.cases, 
          (simp add: cptn.CptnEnv step_e.Env lift_def), 
          (simp add: cptn.CptnEnv step_e.Env_n lift_def))
next                                              
  case CptnComp thus ?case by (simp add: Seqc cptn.CptnComp lift_def)
qed

lemma lift_catch_is_cptn:
assumes cptn_asm:"(Γ,c)∈cptn"
shows "(Γ,map (lift_catch P) c) ∈ cptn"
using cptn_asm
proof (induct)
  case CptnOne thus ?case using cptn.simps by fastforce
next
  case CptnEnv thus ?case  by (cases rule:step_e.cases, 
          (simp add: cptn.CptnEnv step_e.Env lift_catch_def), 
          (simp add: cptn.CptnEnv step_e.Env_n lift_catch_def))
next
  case CptnComp thus ?case by (simp add: Catchc cptn.CptnComp lift_catch_def)
qed

lemma last_lift: "⟦xs≠[]; fst(xs!(length xs - (Suc 0)))=Q⟧ 
 ⟹ fst((map (lift P) xs)!(length (map (lift P) xs)- (Suc 0)))=Seq Q P"
  by (cases "(xs ! (length xs - (Suc 0)))") (simp add:lift_def)

lemma last_lift_catch: "⟦xs≠[]; fst(xs!(length xs - (Suc 0)))=Q⟧ 
 ⟹ fst((map (lift_catch P) xs)!(length (map (lift_catch P) xs)- (Suc 0)))=Catch Q P"
  by (cases "(xs ! (length xs - (Suc 0)))") (simp add:lift_catch_def)

lemma last_fst [rule_format]: "P((a#x)!length x) ⟶ ¬P a ⟶ P (x!(length x - (Suc 0)))" 
  by (induct x) simp_all


lemma last_fst_esp: 
 "fst(((P,s)#xs)!(length xs))=Skip ⟹ P≠Skip ⟹ fst(xs!(length xs - (Suc 0)))=Skip" 
apply(erule last_fst)
apply simp
done

lemma last_snd: "xs≠[] ⟹ 
  snd(((map (lift P) xs))!(length (map (lift P) xs) - (Suc 0)))=snd(xs!(length xs - (Suc 0)))"
  by (cases "(xs ! (length xs - (Suc 0)))") (simp_all add:lift_def)

lemma last_snd_catch: "xs≠[] ⟹ 
  snd(((map (lift_catch P) xs))!(length (map (lift_catch P) xs) - (Suc 0)))=snd(xs!(length xs - (Suc 0)))"
  by (cases "(xs ! (length xs - (Suc 0)))") (simp_all add:lift_catch_def)

lemma Cons_lift: "((Seq P Q), s) # (map (lift Q) xs) = map (lift Q) ((P, s) # xs)"
  by (simp add:lift_def)
thm last_map eq_snd_iff list.inject list.simps(9) last_length
lemma Cons_lift_catch: "((Catch P Q), s) # (map (lift_catch Q) xs) = map (lift_catch Q) ((P, s) # xs)"
  by (simp add:lift_catch_def)

lemma Cons_lift_append: 
  "((Seq P Q), s) # (map (lift Q) xs) @ ys = map (lift Q) ((P, s) # xs)@ ys "
  by (simp add:lift_def)

lemma Cons_lift_catch_append: 
  "((Catch P Q), s) # (map (lift_catch Q) xs) @ ys = map (lift_catch Q) ((P, s) # xs)@ ys "
  by (simp add:lift_catch_def)

lemma lift_nth: "i<length xs ⟹ map (lift Q) xs ! i = lift Q  (xs! i)"
  by (simp add:lift_def)

lemma lift_catch_nth: "i<length xs ⟹ map (lift_catch Q) xs ! i = lift_catch Q  (xs! i)"
  by (simp add:lift_catch_def)
thm list.simps(9) last_length lift_catch_def Cons_lift_catch
lemma snd_lift: "i< length xs ⟹ snd(lift Q (xs ! i))= snd (xs ! i)"
  by (cases "xs!i") (simp add:lift_def)

lemma snd_lift_catch: "i< length xs ⟹ snd(lift_catch Q (xs ! i))= snd (xs ! i)"
  by (cases "xs!i") (simp add:lift_catch_def)

lemma Normal_Normal: 
assumes p1:"(Γ, (P, Normal s) # a # as) ∈ cptn" and       
        p2:"(∃sb. snd (last ((P, Normal s) # a # as))  = Normal sb)"
shows "∃sa. snd a = Normal sa"
proof -
   obtain la1 la2 where last_prod:"last ((P, Normal s)# a#as) = (la1,la2)" by fastforce
   obtain a1 a2 where a_prod:"a=(a1,a2)" by fastforce
   from p1 have clos_p_a:"Γ⊢c (P,Normal s) →ce* (a1, a2)" using a_prod cptn_elim_cases(2)
    proof -
      have f1: "(Γ, (P, Normal s) # (a1, a2) # as) ∈ cptn"
        using a_prod p1 by fastforce
      have "last [(a1, a2)] = (a1, a2)"
        by auto
      thus ?thesis
        using f1 by (metis (no_types) cptn_dest1 cptn_stepconf_rtrancl last_ConsR not_Cons_self2)
    qed  
   then have "Γ⊢c (fst a, snd a) →ce*  (la1,la2)"
   proof -
     from p1 have "(Γ,(a # as)) ∈ cptn" using a_prod cptn_dest by blast
     thus ?thesis by (metis cptn_stepconf_rtrancl last_ConsR last_prod list.distinct(1) prod.collapse) 
   qed 
   then obtain bb where "Normal bb = la2" using last_prod p2 by auto
   thus ?thesis by (metis (no_types) `Γ⊢c (fst a, snd a) →ce* (la1, la2)` steps_ce_not_Normal)
qed
  

lemma lift_P1: 
 assumes map_cptn:"(Γ, map (lift Q) ((P, s) # xs)) ∈ cptn" and
         P_ends:"fst (last ((P, s) # xs)) = Skip"
 shows "(Γ, map (lift Q) ((P, s) # xs) @ [(Q, snd (last ((P, s) # xs)))]) ∈ cptn"
using map_cptn P_ends
proof (induct xs arbitrary: P s) 
  case Nil
  have P0_skips: "P=Skip" using Nil.prems(2) by auto   
  have "(Γ,[(Seq Skip Q, s), (Q, s)]) ∈ cptn"
    by (simp add: cptn.CptnComp SeqSkipc cptn.CptnOne)
  then show ?case using P0_skips by (simp add: lift_def)
next
  case (Cons a xs)
  have "(Γ, map (lift Q) ((P, s) # a # xs)) ∈ cptn"
    using Cons.prems(1) by blast  
  have "fst (last ( a # xs)) = Skip" using Cons.prems(2) by auto
  also have seq_PQ:"(Γ,(Seq P Q,s) # (map (lift Q) (a#xs))) ∈ cptn"
    by (metis Cons.prems(1) Cons_lift)
  then have "(Γ,(map (lift Q) (a#xs))) ∈ cptn"
    proof -
      assume a1:"(Γ, (Seq P Q, s) # map (lift Q) (a # xs)) ∈ cptn"            
      then obtain a1 a2 xs1 where a2: "map (lift Q) (a#xs) = ((a1,a2)#xs1)" by fastforce 
      thus ?thesis  using cptn_dest using seq_PQ by auto 
    qed 
  then have "(Γ, map (lift Q) (a#xs) @ [(Q, snd (last ((a#xs))))]) ∈ cptn" 
   by (metis Cons.hyps(1) calculation prod.collapse)   
  then have t1:"(Γ, (Seq (fst a) Q, (snd a))#map (lift Q) xs @ [(Q, snd (last ((P, s)#(a#xs))))]) ∈ cptn"
   by (simp add: Cons_lift_append)
  then have "(Γ,(Seq P Q,s) # (Seq (fst a) Q, (snd a))#map (lift Q) xs)∈ cptn"
   using seq_PQ by (simp add: Cons_lift)  
  then have t2: "(Γ,(Seq P Q,s) # [(Seq (fst a) Q, (snd a))]) ∈ cptn"
   using cptn_dest1 by blast
  then have"((Seq P Q,s) # [(Seq (fst a) Q, (snd a))])!length [(Seq (fst a) Q, (snd a))] = (Seq (fst a) Q, (snd a))"
   by auto  
  then have "(Γ,(Seq P Q,s) # [(Seq (fst a) Q, (snd a))]@map (lift Q) xs @ [(Q, snd (last ((P, s)#(a#xs))))])∈ cptn"
   using cptn_append_is_cptn t1 t2 by blast   
  then have "(Γ, map (lift Q) ((P,s)#(fst a, snd a)#xs) @[(Q, snd (last ((P, s)#(a#xs))))])∈cptn"
   using Cons_lift_append append_Cons append_Nil by metis
  thus ?case by auto    
qed


lemma lift_catch_P1: 
 assumes map_cptn:"(Γ, map (lift_catch Q) ((P, Normal s) # xs)) ∈ cptn" and
         P_ends:"fst (last ((P, Normal s) # xs)) = Throw" and
         P_ends_normal:"∃p. snd(last ((P, Normal s) # xs)) = Normal p"
 shows "(Γ, map (lift_catch Q) ((P, Normal s) # xs) @ [(Q, snd (last ((P, Normal s) # xs)))]) ∈ cptn"
using map_cptn P_ends P_ends_normal
proof (induct xs arbitrary: P s) 
  case Nil
  have P0_skips: "P=Throw" using Nil.prems(2) by auto   
  have "(Γ,[(Catch Throw Q, Normal s), (Q, Normal s)]) ∈ cptn"
    by (simp add: cptn.CptnComp CatchThrowc cptn.CptnOne)
  then show ?case using P0_skips by (simp add: lift_catch_def)
next
  case (Cons a xs) 
  have s1:"(Γ, map (lift_catch Q) ((P, Normal s) # a # xs)) ∈ cptn"
    using Cons.prems(1) by blast 
  have s2:"fst (last ( a # xs)) = Throw" using Cons.prems(2) by auto
  then obtain p where s3:"snd(last (a #xs)) = Normal p" using Cons.prems(3) by auto
  also have seq_PQ:"(Γ,(Catch P Q,Normal s) # (map (lift_catch Q) (a#xs))) ∈ cptn"
    by (metis Cons.prems(1) Cons_lift_catch) thm Cons.hyps
  then have axs_in_cptn:"(Γ,(map (lift_catch Q) (a#xs))) ∈ cptn"
    proof -
      assume a1:"(Γ, (Catch P Q, Normal s) # map (lift_catch Q) (a # xs)) ∈ cptn"            
      then obtain a1 a2 xs1 where a2: "map (lift_catch Q) (a#xs) = ((a1,a2)#xs1)" by fastforce 
      thus ?thesis  using cptn_dest using seq_PQ by auto 
    qed 
  then have "(Γ, map (lift_catch Q) (a#xs) @ [(Q, snd (last ((a#xs))))]) ∈ cptn" 
    proof (cases "xs=[]")
      case True thus ?thesis using s2 s3 axs_in_cptn by (metis Cons.hyps eq_snd_iff last_ConsL)
    next
      case False            
      from seq_PQ have seq:"(Γ,(Catch P Q,Normal s) # (Catch (fst a) Q,snd a)#map (lift_catch Q) xs)∈ cptn"
        by (simp add: Cons_lift_catch)                         
      obtain cf sf where last_map_axs:"(cf,sf)=last (map (lift_catch Q) (a#xs))" using prod.collapse by blast
      have "∀p ps. (ps=[] ∧ last [p] = p) ∨ (ps≠[] ∧ last (p#ps) = last ps)" by simp              
      then have tranclos:"Γ⊢c (Catch P Q,Normal s) →ce* (Catch (fst a) Q,snd a)" using Cons_lift_catch
            by (metis (no_types) cptn_dest1 cptn_stepc_rtrancl not_Cons_self2 seq)
      have tranclos_a:"Γ⊢c (Catch (fst a) Q,snd a) →ce* (cf,sf)"
           by (metis Cons_lift_catch axs_in_cptn cptn_stepc_rtrancl last_map_axs prod.collapse)      
      have snd_last:"snd (last (map (lift_catch Q) (a#xs))) = snd (last (a #xs))"
      proof - 
         have eqslist:"snd(((map (lift_catch Q) (a#xs)))!(length (map (lift_catch Q) xs)))= snd((a#xs)!(length xs))" 
           using last_snd_catch by fastforce  
         have "(lift_catch Q a)#(map (lift_catch Q) xs) = (map (lift_catch Q) (a#xs))" by auto
         then have "(map (lift_catch Q) (a#xs))!(length (map (lift_catch Q) xs)) = last (map (lift_catch Q) (a#xs))"
           using last_length [of "(lift_catch Q a)" "(map (lift_catch Q) xs)"] by auto
         thus ?thesis using eqslist by (simp add:last_length)
      qed
      then obtain p1 where  "(snd a) = Normal p1"
         by (metis tranclos_a last_map_axs s3 snd_conv step_ce_normal_to_normal tranclos)   
      moreover obtain a1 a2 where aeq:"a = (a1,a2)" by fastforce 
      moreover have "fst (last ((a1,a2) # xs)) = Throw" using s2 False by auto  
      moreover have "(Γ, map (lift_catch Q) ((a1,a2) # xs)) ∈ cptn" using aeq axs_in_cptn False by auto  
      moreover have "∃p. snd (last ((a1,a2) # xs)) = Normal p" using s3 aeq by auto
      moreover have "a2 = Normal p1" using aeq calculation(1) by auto 
      ultimately have "(Γ, map (lift_catch Q) ((a1,a2) # xs) @
                           [(Q, snd (last ((a1,a2) # xs)))])∈ cptn" 
                 using Cons.hyps aeq by blast
      thus ?thesis using aeq by force 
    qed      
  then have t1:"(Γ, (Catch (fst a) Q, (snd a))#map (lift_catch Q) xs @ [(Q, snd (last ((P, Normal s)#(a#xs))))]) ∈ cptn"
   by (simp add: Cons_lift_catch_append)
  then have "(Γ,(Catch P Q,Normal s) # (Catch (fst a) Q, (snd a))#map (lift_catch Q) xs)∈ cptn"
   using seq_PQ by (simp add: Cons_lift_catch)  
  then have t2: "(Γ,(Catch P Q,Normal s) # [(Catch (fst a) Q, (snd a))]) ∈ cptn"
   using cptn_dest1 by blast
  then have"((Catch P Q,Normal s) # [(Catch (fst a) Q, (snd a))])!length [(Catch (fst a) Q, (snd a))] = (Catch (fst a) Q, (snd a))"
   by auto  
  then have "(Γ,(Catch P Q,Normal s) # [(Catch (fst a) Q, (snd a))]@map (lift_catch Q) xs @ [(Q, snd (last ((P, Normal s)#(a#xs))))])∈ cptn"
   using cptn_append_is_cptn t1 t2 by blast
  then have "(Γ, map (lift_catch Q) ((P,Normal s)#(fst a, snd a)#xs) @[(Q, snd (last ((P,Normal s)#(a#xs))))])∈cptn"
   using Cons_lift_catch_append append_Cons append_Nil by metis
  thus ?case by auto    
qed

lemma seq2:
assumes 
    p1:"(Γ, (P0, s) # xs) ∈ cptn_mod" and
    p2:"(Γ, (P0, s) # xs) ∈ cptn" and
    p3:"fst (last ((P0, s) # xs)) = Skip" and
    p4:"(Γ, (P1, snd (last ((P0, s) # xs))) # ys) ∈ cptn_mod" and
    p5:"(Γ, (P1, snd (last ((P0, s) # xs))) # ys) ∈ cptn" and
    p6:"zs = map (lift P1) xs @ (P1, snd (last ((P0, s) # xs))) # ys"
shows "(Γ, (Seq P0 P1, s) # zs) ∈ cptn"
using p1 p2 p3 p4 p5 p6
proof -
have last_skip:"fst (last ((P0, s) # xs)) = Skip" using p3 by blast 
  have "(Γ, (map (lift P1) ((P0, s) # xs))@(P1, snd (last ((P0, s) # xs))) # ys) ∈ cptn"
  proof -
    have "(Γ,map (lift P1) ((P0, s) #xs)) ∈ cptn"
      using p2 lift_is_cptn by blast 
    then have "(Γ,map (lift P1) ((P0, s) #xs)@[(P1, snd (last ((P0, s) # xs)))]) ∈ cptn" 
      using last_skip lift_P1 by blast 
    then have "(Γ,(Seq P0 P1, s) # map (lift P1) xs@[(P1, snd (last ((P0, s) # xs)))]) ∈ cptn"
         by (simp add: Cons_lift_append)
    moreover have "last ((Seq P0 P1, s) # map (lift P1) xs @[(P1, snd (last ((P0, s) # xs)))]) = (P1, snd (last ((P0, s) # xs)))" 
      by auto  
    moreover have "last ((Seq P0 P1, s) # map (lift P1) xs @[(P1, snd (last ((P0, s) # xs)))]) =
                   ((Seq P0 P1, s) # map (lift P1) xs @[(P1, snd (last ((P0, s) # xs)))])!length (map (lift P1) xs @[(P1, snd (last ((P0, s) # xs)))])" 
      by (metis last_length)             
    ultimately have "(Γ, (Seq P0 P1, s) # map (lift P1) xs @ (P1, snd (last ((P0, s) # xs))) # ys)∈ cptn" 
      using cptn_append_is_cptn p5 by fastforce
    thus ?thesis by (simp add: Cons_lift_append)
  qed 
  thus ?thesis  
    by (simp add: Cons_lift_append p6)
qed

lemma seq3:
assumes
    p1:"(Γ, (P0, Normal s) # xs) ∈ cptn_mod" and
    p2:"(Γ, (P0, Normal s) # xs) ∈ cptn" and
    p3:"fst (last ((P0, Normal s) # xs)) = Throw" and
    p4:"snd (last ((P0, Normal s) # xs)) = Normal s'" and
    p5:"(Γ,(Throw,Normal s')#ys) ∈ cptn_mod" and
    p6:"(Γ,(Throw,Normal s')#ys) ∈ cptn" and
    p7:"zs = map (lift P1) xs @((Throw,Normal s')#ys)" 
shows "(Γ, (Seq P0 P1, Normal s) # zs) ∈ cptn"
using p1 p2 p3 p4 p5 p6 p7
proof (induct xs arbitrary: zs P0 s) 
  case Nil thus ?case using SeqThrowc cptn.simps by fastforce
next
  case (Cons a as)
  then obtain sa where "snd a = Normal sa" by (meson Normal_Normal)
  obtain a1 a2 where a_prod:"a=(a1,a2)" by fastforce
  obtain la1 la2 where last_prod:"last (a#as) = (la1,la2)" by fastforce 
  then have lasst_aas_last: "last (a#as) = (last ((P0, Normal s) # a # as))" by auto    
  then have "la1 = Throw" using Cons.prems(3) last_prod by force
  have "la2 = Normal s'" using Cons.prems(4) last_prod lasst_aas_last by force
  have f1: "(Γ, (a1, a2) # as) ∈ cptn"
    using Cons.prems(2) a_prod cptn_dest by blast
  have f2: "Normal sa = a2"
    using `snd a = Normal sa` a_prod by force
  have "(Γ, a # as) ∈ cptn_mod"
    using f1 a_prod cptn_onlyif_cptn_mod by blast
  then have hyp:"(Γ, (Seq a1 P1, Normal sa) # 
            map (lift P1) as @ ((Throw,Normal s')#ys)) ∈ cptn"
        using Cons.hyps Cons.prems(3) Cons.prems(4)  Cons.prems(5) Cons.prems(6) a_prod f1 f2 by fastforce
  thus ?case
  proof -
    have "(Seq a1 P1, a2) # map (lift P1) as @((Throw,Normal s')#ys) = zs"
      by (simp add: Cons.prems(7) Cons_lift_append a_prod)
    thus ?thesis
      by (metis (no_types, lifting) Cons.prems(2) Seqc a_prod cptn.CptnComp cptn.CptnEnv Env cptn_elim_cases(2) f2 hyp)     
  qed 
qed

lemma cptn_if_cptn_mod: 
assumes cptn_mod_asm:"(Γ,c) ∈ cptn_mod"
shows  "(Γ,c) ∈ cptn"
using cptn_mod_asm
proof (induct)
  case (CptnModOne) thus ?case using cptn.CptnOne by blast
next
  case CptnModSkip thus ?case by (simp add: cptn.CptnComp) 
next
  case CptnModThrow thus ?case by (simp add: cptn.CptnComp) 
next 
  case CptnModCondT thus ?case by (simp add: CondTruec cptn.CptnComp) 
next
  case CptnModCondF thus ?case by (simp add: CondFalsec cptn.CptnComp)
next
  case (CptnModSeq1 Γ P0 s xs zs P1) 
  have "(Γ, map (lift P1) ((P0, s) # xs)) ∈ cptn"
    using CptnModSeq1.hyps(2) lift_is_cptn by blast
  thus ?case by (simp add: Cons_lift CptnModSeq1.hyps(3))
next
  case (CptnModSeq2 Γ P0 s xs P1 ys zs)   
  thus ?case by (simp add:seq2)
next
  case (CptnModSeq3 Γ P0 s xs s' zs P1) 
  thus ?case by (simp add: seq3)
next
  case (CptnModWhile1  Γ P s xs b zs) thus ?case by (metis Cons_lift WhileTruec cptn.CptnComp lift_is_cptn)
next 
  case (CptnModWhile2  Γ P s xs b zs ys)    
  then have "(Γ, (Seq P (While b P), Normal s) # zs) ∈ cptn" 
    by (simp add:seq2)  
  then have "Γ⊢c(While b P,Normal s) → (Seq P (While b P),Normal s)" 
    by (simp add: CptnModWhile2.hyps(4) WhileTruec)
  thus ?case
   by (simp add: `(Γ, (Seq P (While b P), Normal s) # zs) ∈ cptn` cptn.CptnComp) 
next
  case (CptnModWhile3  Γ P s xs b s' ys zs) 
  then have "(Γ,(Seq P (While b P), Normal s) # zs) ∈ cptn"
     by (simp add: seq3)     
  then have "Γ⊢c(While b P,Normal s) → (Seq P (While b P),Normal s)" by (simp add: CptnModWhile3.hyps(4) WhileTruec)
  thus ?case by (simp add: `(Γ, (Seq P (While b P), Normal s) # zs) ∈ cptn` cptn.CptnComp)
next 
  case (CptnModCall Γ bdy s ys p) thus ?case by (simp add: Callc cptn.CptnComp) 
next
  case (CptnModDynCom Γ c s ys) thus ?case by (simp add: DynComc cptn.CptnComp)
next
  case (CptnModGuard Γ c s ys g f) thus ?case by (simp add: Guardc cptn.CptnComp)
next
  case (CptnModCatch1 Γ P0 s xs zs P1)
  have "(Γ, map (lift_catch P1) ((P0, s) # xs)) ∈ cptn"
    using CptnModCatch1.hyps(2) lift_catch_is_cptn by blast
  thus ?case by (simp add: Cons_lift_catch CptnModCatch1.hyps(3))
next
  case (CptnModCatch2 Γ P0 s xs ys zs P1)
  thus ?case
  proof (induct xs arbitrary: zs P0 s) 
    case Nil thus ?case using CatchSkipc cptn.simps by fastforce
  next
    case (Cons a as)
    then obtain sa where "snd a = sa" by auto
    then obtain a1 a2 where a_prod:"a=(a1,a2)" and sa_a2: "a2 =sa" 
           by fastforce
    obtain la1 la2 where last_prod:"last (a#as) = (la1,la2)" by fastforce 
    then have lasst_aas_last: "last (a#as) = (last ((P0, s) # a # as))" by auto    
    then have "la1 = Skip" using Cons.prems(3) last_prod by force
    have f1: "(Γ, (a1, a2) # as) ∈ cptn"
      using Cons.prems(2) a_prod cptn_dest by blast
    have "(Γ, a # as) ∈ cptn_mod"
      using f1 a_prod cptn_onlyif_cptn_mod by blast
    then have hyp:"(Γ, (Catch a1 P1, a2) # 
              map (lift_catch P1) as @ ((Skip, la2)#ys)) ∈ cptn"
          using Cons.hyps Cons.prems a_prod f1 last_prod by fastforce
    thus ?case
    proof -
      have f1:"(Catch a1 P1, a2) # map (lift_catch P1) as @ ((Skip, la2)#ys) = zs"
        using Cons.prems(4) Cons_lift_catch_append a_prod last_prod by (simp add: Cons.prems(6))         
      have "(Γ, map (lift_catch P1) ((P0, s) # a # as)) ∈ cptn"
       using Cons.prems(2) lift_catch_is_cptn by blast
      hence "(Γ, (LanguageCon.com.Catch P0 P1, s) # (LanguageCon.com.Catch a1 P1, a2) # map (lift_catch P1) as) ∈ cptn"
        by (metis (no_types) Cons_lift_catch a_prod)
      hence "(Γ, (LanguageCon.com.Catch P0 P1, s) # zs) ∈ cptn ∨ (Γ, (LanguageCon.com.Catch P0 P1, s) # (LanguageCon.com.Catch a1 P1, a2) # map (lift_catch P1) as) ∈ cptn ∧ (¬ Γ⊢c (LanguageCon.com.Catch P0 P1, s) →e (LanguageCon.com.Catch P0 P1, a2) ∨ (Γ, (LanguageCon.com.Catch P0 P1, a2) # map (lift_catch P1) as) ∉ cptn ∨ LanguageCon.com.Catch a1 P1 ≠ LanguageCon.com.Catch P0 P1)"
        using f1 cptn.CptnEnv hyp by blast
      thus ?thesis
       by (metis (no_types) f1 cptn.CptnComp cptn_elim_cases(2) hyp)
     qed
   qed
next 
  case (CptnModCatch3  Γ P0 s xs s' P1 ys zs)
  thus ?case
  proof (induct xs arbitrary: zs P0 s) 
    case Nil thus ?case using CatchThrowc cptn.simps by fastforce
  next
    case (Cons a as)
    then obtain sa where "snd a = Normal sa" by (meson Normal_Normal)
    obtain a1 a2 where a_prod:"a=(a1,a2)" by fastforce
    obtain la1 la2 where last_prod:"last (a#as) = (la1,la2)" by fastforce 
    then have lasst_aas_last: "last (a#as) = (last ((P0, Normal s) # a # as))" by auto    
    then have "la1 = Throw" using Cons.prems(3) last_prod by force
    have "la2 = Normal s'" using Cons.prems(4) last_prod lasst_aas_last by force
    have f1: "(Γ, (a1, a2) # as) ∈ cptn"
      using Cons.prems(2) a_prod cptn_dest by blast
    have f2: "Normal sa = a2"
      using `snd a = Normal sa` a_prod by force
    have "(Γ, a # as) ∈ cptn_mod"
      using f1 a_prod cptn_onlyif_cptn_mod by blast
    then have hyp:"(Γ, (Catch a1 P1, Normal sa) # 
              map (lift_catch P1) as @ (P1, snd (last ((a1, Normal sa) # as))) # ys) ∈ cptn"
          using Cons.hyps Cons.prems a_prod f1 f2 by auto
    thus ?case
    proof -
      have "Γ⊢c (P0, Normal s) →e (P0, a2)"
        by (fastforce intro: step_e.intros)
      then have transit:"Γ⊢c(P0,Normal s) →ce (a1,Normal sa)" 
            by (metis (no_types) Cons.prems(2) a_prod c_step cptn_elim_cases(2) e_step f2)
      then have transit_catch:"Γ⊢c(Catch P0 P1,Normal s) →ce (Catch a1 P1,Normal sa)"
            by (metis (no_types) Catchc c_step e_step env_c_c' step_ce_elim_cases step_e.intros(1)) 
      have "(Catch a1 P1, a2) # map (lift_catch P1) as @ (P1, la2) # ys = zs"
        using Cons.prems Cons_lift_catch_append a_prod last_prod by auto        
      have "a=(a1, Normal sa)" using a_prod f2 by auto
      have "snd (last ((a1, Normal sa) # as)) = Normal s'"
          using `a = (a1, Normal sa)` `snd (last ((P0, Normal s) # a # as)) = Normal s'` lasst_aas_last by fastforce
      hence f1: "snd (last ((a1, Normal sa) # as)) = la2"
          using `la2 = Normal s'` by blast
      have "Γ⊢c (LanguageCon.com.Catch P0 P1, Normal s) →ce (LanguageCon.com.Catch a1 P1, a2)"
          using f2 transit_catch by blast
      thus ?thesis
        using f1 `(LanguageCon.com.Catch a1 P1, a2) # map (lift_catch P1) as @ (P1, la2) # ys = zs`  
              cptn.CptnComp cptn.CptnEnv f2 hyp not_eq_not_env step_ce_not_step_e_step_c 
        by metis
    qed
  qed
next 
  case (CptnModEnv) thus ?case by (simp add: cptn.CptnEnv)
qed

lemma cptn_eq_cptn_mod: 
shows  "(x ∈cptn_mod)  = (x∈cptn)"
by (cases x, auto simp add: cptn_if_cptn_mod cptn_onlyif_cptn_mod)

lemma cptn_eq_cptn_mod_set: 
shows  "cptn_mod  = cptn"
by (auto simp add: cptn_if_cptn_mod cptn_onlyif_cptn_mod)

subsection ‹Computational modular semantic for nested calls›
inductive_set cptn_mod_nest_call :: "(nat×('s,'p,'f,'e) confs) set"
where
  CptnModNestOne: "(n,Γ,[(P, s)]) ∈ cptn_mod_nest_call"
| CptnModNestEnv: "⟦Γ⊢c(P,s) →e (P,t);(n,Γ,(P, t)#xs) ∈ cptn_mod_nest_call⟧ ⟹ 
                     (n,Γ,(P, s)#(P, t)#xs) ∈ cptn_mod_nest_call"
| CptnModNestSkip: "⟦Γ⊢c(P,s) → (Skip,t); redex P = P; 
                     ∀f. ((∃sn. s = Normal sn) ∧ (Γ f) = Some Skip ⟶ P  ≠ Call f  );
                (n,Γ,(Skip, t)#xs) ∈ cptn_mod_nest_call ⟧ ⟹ 
                (n,Γ,(P,s)#(Skip, t)#xs) ∈cptn_mod_nest_call"

| CptnModNestThrow: "⟦Γ⊢c(P,s) → (Throw,t); redex P = P; 
                     ∀f. ((∃sn. s = Normal sn) ∧ (Γ f) = Some Throw ⟶ P  ≠ Call f  );
                      (n,Γ,(Throw, t)#xs) ∈ cptn_mod_nest_call ⟧ ⟹ 
                      (n,Γ,(P,s)#(Throw, t)#xs) ∈cptn_mod_nest_call"

| CptnModNestCondT: "⟦(n,Γ,(P0, Normal s)#ys) ∈ cptn_mod_nest_call; s ∈ b ⟧ ⟹ 
                    (n,Γ,((Cond b P0 P1), Normal s)#(P0, Normal s)#ys) ∈ cptn_mod_nest_call"

| CptnModNestCondF: "⟦(n,Γ,(P1, Normal s)#ys) ∈ cptn_mod_nest_call; s ∉ b ⟧ ⟹ 
                     (n,Γ,((Cond b P0 P1), Normal s)#(P1, Normal s)#ys) ∈ cptn_mod_nest_call"

| CptnModNestSeq1: 
  "⟦(n,Γ,(P0, s)#xs) ∈ cptn_mod_nest_call; zs=map (lift P1) xs ⟧ ⟹ 
   (n,Γ,((Seq P0 P1), s)#zs) ∈ cptn_mod_nest_call"

| CptnModNestSeq2: 
  "⟦(n,Γ, (P0, s)#xs) ∈ cptn_mod_nest_call; fst(last ((P0, s)#xs)) = Skip;
    (n,Γ,(P1, snd(last ((P0, s)#xs)))#ys) ∈ cptn_mod_nest_call;
    zs=(map (lift P1) xs)@((P1, snd(last ((P0, s)#xs)))#ys) ⟧ ⟹ 
   (n,Γ,((Seq P0 P1), s)#zs) ∈ cptn_mod_nest_call"

| CptnModNestSeq3: 
  "⟦(n,Γ, (P0, Normal s)#xs) ∈ cptn_mod_nest_call; 
    fst(last ((P0, Normal s)#xs)) = Throw;
    snd(last ((P0, Normal s)#xs)) = Normal s'; 
    (n,Γ,(Throw,Normal s')#ys) ∈ cptn_mod_nest_call; 
    zs=(map (lift P1) xs)@((Throw,Normal s')#ys) ⟧ ⟹
   (n,Γ,((Seq P0 P1), Normal s)#zs) ∈ cptn_mod_nest_call"

| CptnModNestWhile1: 
  "⟦(n,Γ, (P, Normal s)#xs) ∈ cptn_mod_nest_call; s ∈ b; 
    zs=map (lift (While b P)) xs ⟧ ⟹ 
    (n,Γ, ((While b P), Normal s)#
      ((Seq P (While b P)),Normal s)#zs) ∈ cptn_mod_nest_call"

| CptnModNestWhile2: 
  "⟦ (n,Γ, (P, Normal s)#xs) ∈ cptn_mod_nest_call; 
     fst(last ((P, Normal s)#xs))=Skip; s ∈ b; 
     zs=(map (lift (While b P)) xs)@
      (While b P, snd(last ((P, Normal s)#xs)))#ys; 
      (n,Γ,(While b P, snd(last ((P, Normal s)#xs)))#ys) ∈ 
        cptn_mod_nest_call⟧ ⟹ 
   (n,Γ,(While b P, Normal s)#
     (Seq P (While b P), Normal s)#zs) ∈ cptn_mod_nest_call"

| CptnModNestWhile3: 
  "⟦ (n,Γ, (P, Normal s)#xs) ∈ cptn_mod_nest_call; 
     fst(last ((P, Normal s)#xs))=Throw; s ∈ b;
     snd(last ((P, Normal s)#xs)) = Normal s'; 
    (n,Γ,(Throw,Normal s')#ys) ∈ cptn_mod_nest_call;  
     zs=(map (lift (While b P)) xs)@((Throw,Normal s')#ys)⟧ ⟹ 
   (n,Γ,(While b P, Normal s)#
     (Seq P (While b P), Normal s)#zs) ∈ cptn_mod_nest_call"

| CptnModNestCall: "⟦(n,Γ,(bdy, Normal s)#ys) ∈ cptn_mod_nest_call;Γ p = Some bdy; bdy≠Call p ⟧ ⟹ 
                 (Suc n, Γ,((Call p), Normal s)#(bdy, Normal s)#ys) ∈ cptn_mod_nest_call" 

| CptnModNestDynCom: "⟦(n,Γ,(c s, Normal s)#ys) ∈ cptn_mod_nest_call ⟧ ⟹ 
                 (n,Γ,(DynCom c, Normal s)#(c s, Normal s)#ys) ∈ cptn_mod_nest_call"

| CptnModNestGuard: "⟦(n,Γ,(c, Normal s)#ys) ∈ cptn_mod_nest_call; s ∈ g ⟧ ⟹ 
                  (n,Γ,(Guard f g c, Normal s)#(c, Normal s)#ys) ∈ cptn_mod_nest_call"

| CptnModNestCatch1: "⟦(n,Γ,(P0, s)#xs) ∈ cptn_mod_nest_call; zs=map (lift_catch P1) xs ⟧
                 ⟹ (n,Γ,((Catch P0 P1), s)#zs) ∈ cptn_mod_nest_call"

| CptnModNestCatch2: 
  "⟦(n,Γ, (P0, s)#xs) ∈ cptn_mod_nest_call; fst(last ((P0, s)#xs)) = Skip; 
    (n,Γ,(Skip,snd(last ((P0, s)#xs)))#ys) ∈ cptn_mod_nest_call;  
    zs=(map (lift_catch P1) xs)@((Skip,snd(last ((P0, s)#xs)))#ys) ⟧ ⟹ 
   (n,Γ,((Catch P0 P1), s)#zs) ∈ cptn_mod_nest_call"

| CptnModNestCatch3: 
  "⟦(n,Γ, (P0, Normal s)#xs) ∈ cptn_mod_nest_call; fst(last ((P0, Normal s)#xs)) = Throw; 
  snd(last ((P0, Normal s)#xs)) = Normal s';
  (n,Γ,(P1, snd(last ((P0, Normal s)#xs)))#ys) ∈ cptn_mod_nest_call; 
  zs=(map (lift_catch P1) xs)@((P1, snd(last ((P0, Normal s)#xs)))#ys) ⟧ ⟹ 
   (n,Γ,((Catch P0 P1), Normal s)#zs) ∈ cptn_mod_nest_call"

lemmas CptnMod_nest_call_induct = cptn_mod_nest_call.induct [of _ _ "[(c,s)]", split_format (complete), case_names
CptnModOne CptnModEnv CptnModSkip CptnModThrow CptnModCondT CptnModCondF 
CptnModSeq1 CptnModSeq2 CptnModSeq3 CptnModSeq4 CptnModWhile1 CptnModWhile2 CptnModWhile3 CptnModCall CptnModDynCom CptnModGuard 
CptnModCatch1 CptnModCatch2 CptnModCatch3, induct set]

inductive_cases CptnModNest_elim_cases [cases set]:
"(n,Γ,(Skip, s)#u#xs) ∈ cptn_mod_nest_call"
"(n,Γ,(Guard f g c, s)#u#xs) ∈ cptn_mod_nest_call"
"(n,Γ,(Basic f e, s)#u#xs) ∈ cptn_mod_nest_call"
"(n,Γ,(Spec r e, s)#u#xs) ∈ cptn_mod_nest_call"
"(n,Γ,(Seq c1 c2, s)#u#xs) ∈ cptn_mod_nest_call"
"(n,Γ,(Cond b c1 c2, s)#u#xs) ∈ cptn_mod_nest_call"
"(n,Γ,(Await b c2 e, s)#u#xs) ∈ cptn_mod_nest_call"
"(n,Γ,(Call p, s)#u#xs) ∈ cptn_mod_nest_call"
"(n,Γ,(DynCom c,s)#u#xs) ∈ cptn_mod_nest_call"
"(n,Γ,(Throw,s)#u#xs) ∈ cptn_mod_nest_call"
"(n,Γ,(Catch c1 c2,s)#u#xs) ∈ cptn_mod_nest_call"

inductive_cases stepc_elim_cases_Seq_Seq':
"Γ⊢c(Seq c1 c2,s) → (Seq c1' c2',s')" 

inductive_cases stepc_elim_cases_Catch_Catch':
"Γ⊢c(Catch c1 c2,s) → (Catch c1' c2',s')" 

inductive_cases CptnModNest_same_elim_cases [cases set]:
"(n,Γ,(u, s)#(u,t)#xs) ∈ cptn_mod_nest_call"

inductive_cases CptnModNest_elim_cases_Stuck [cases set]:
"(n,Γ,(P, Stuck)#(Skip, s)#xs) ∈ cptn_mod_nest_call"

inductive_cases CptnModNest_elim_cases_Fault [cases set]:
"(n,Γ,(P, Fault f)#(Skip, s)#xs) ∈ cptn_mod_nest_call"

inductive_cases CptnModNest_elim_cases_Abrupt [cases set]:
"(n,Γ,(P, Abrupt as)#(Skip, s)#xs) ∈ cptn_mod_nest_call"

inductive_cases  CptnModNest_elim_cases_Call_Stuck [cases set]:
"(n,Γ,(Call p, s)#(Skip, Stuck)#xs) ∈ cptn_mod_nest_call"

inductive_cases  CptnModNest_elim_cases_Call [cases set]:
"(0, Γ,((Call p), Normal s)#(bdy, Normal s)#ys) ∈ cptn_mod_nest_call"

inductive_cases  CptnEmpty [cases set]:
"(n, Γ,[]) ∈ cptn_mod_nest_call"

inductive_cases  CptnModNest_elim_cases_Call_normal [cases set]:
"(Suc n, Γ,((Call p), Normal s)#(bdy, Normal s)#ys) ∈ cptn_mod_nest_call"

lemma cptn_mod_nest_mono1: "(n,Γ,cfs) ∈  cptn_mod_nest_call  ⟹ (Suc n,Γ,cfs)∈ cptn_mod_nest_call"
proof (induct rule:cptn_mod_nest_call.induct)
  case (CptnModNestOne) thus ?case using cptn_mod_nest_call.CptnModNestOne by auto
next
  case (CptnModNestEnv) thus ?case using cptn_mod_nest_call.CptnModNestEnv by fastforce
next
   case (CptnModNestSkip) thus ?case using cptn_mod_nest_call.CptnModNestSkip by fastforce
next
   case (CptnModNestThrow) thus ?case using cptn_mod_nest_call.intros(4)  by fastforce
next
   case (CptnModNestCondT n) thus ?case 
     using cptn_mod_nest_call.CptnModNestCondT[of "Suc n"] by fastforce
next
  case (CptnModNestCondF n) thus ?case 
    using cptn_mod_nest_call.CptnModNestCondF[of "Suc n"] by fastforce
next
  case (CptnModNestSeq1 n) thus ?case 
    using cptn_mod_nest_call.CptnModNestSeq1[of "Suc n"] by fastforce
next
  case (CptnModNestSeq2 n) thus ?case 
     using cptn_mod_nest_call.CptnModNestSeq2[of "Suc n"] by fastforce
next
  case (CptnModNestSeq3 n) thus ?case 
     using cptn_mod_nest_call.CptnModNestSeq3[of "Suc n"] by fastforce
next
  case (CptnModNestWhile1 n) thus ?case 
     using cptn_mod_nest_call.CptnModNestWhile1[of "Suc n"] by fastforce
next
  case (CptnModNestWhile2 n) thus ?case 
     using cptn_mod_nest_call.CptnModNestWhile2[of "Suc n"] by fastforce
next
  case (CptnModNestWhile3 n) thus ?case 
     using cptn_mod_nest_call.CptnModNestWhile3[of "Suc n"] by fastforce
next
 case (CptnModNestCall) thus ?case 
     using cptn_mod_nest_call.CptnModNestCall by fastforce
next 
 case (CptnModNestDynCom) thus ?case 
     using cptn_mod_nest_call.CptnModNestDynCom by fastforce
next
 case (CptnModNestGuard n) thus ?case 
     using cptn_mod_nest_call.CptnModNestGuard[of "Suc n"] by fastforce
next
 case (CptnModNestCatch1 n) thus ?case 
     using cptn_mod_nest_call.CptnModNestCatch1[of "Suc n"] by fastforce
next
 case (CptnModNestCatch2 n) thus ?case 
     using cptn_mod_nest_call.CptnModNestCatch2[of "Suc n"] by fastforce
next
 case (CptnModNestCatch3 n) thus ?case 
     using cptn_mod_nest_call.CptnModNestCatch3[of "Suc n"] by fastforce
qed

lemma cptn_mod_nest_mono2: 
  "(n,Γ,cfs) ∈  cptn_mod_nest_call  ⟹ m>n ⟹
   (m,Γ,cfs)∈ cptn_mod_nest_call"
proof (induct "m-n" arbitrary: m n)
  case 0 thus ?case by auto
next
  case (Suc k) 
  have "m - Suc n = k"
    using Suc.hyps(2) Suc.prems(2) Suc_diff_Suc Suc_inject by presburger
  then show ?case
    using Suc.hyps(1) Suc.prems(1) Suc.prems(2) cptn_mod_nest_mono1 less_Suc_eq by blast   
qed

lemma cptn_mod_nest_mono: 
  "(n,Γ,cfs) ∈  cptn_mod_nest_call  ⟹ m≥n ⟹
   (m,Γ,cfs)∈ cptn_mod_nest_call"
proof (cases "n=m")
  assume "(n, Γ, cfs) ∈ cptn_mod_nest_call" and  
          "n = m"  thus ?thesis by auto
next
  assume "(n, Γ, cfs) ∈ cptn_mod_nest_call" and  
         "n≤m" and
         "n ≠ m"  
   thus ?thesis by (auto simp add: cptn_mod_nest_mono2)
qed

subsection ‹Lemmas on normalization›

(* lemma step_sequence_flatten:
  assumes exec: "Γ⊢c(P,s) → (Q,t)" 
  shows "Γ⊢c(sequence Seq (flatten P),s) → (sequence Seq (flatten Q),t)"
using exec
proof (induct rule: stepc_induct)
  case (Guardc s g f c) thus ?case using stepc.Guardc
  case (Seqc c1 s c2 s' c2')
  then have " Γ⊢c (Seq (LanguageCon.sequence LanguageCon.com.Seq (LanguageCon.flatten c1)) c2', s) →
                   (Seq (LanguageCon.sequence LanguageCon.com.Seq (LanguageCon.flatten c2)) c2', s') " 
    using stepc.Seqc by fastforce    
  thus ?case sorry
qed(auto intro:stepc.intros)+

lemma normalice_step:
  assumes exec:"Γ⊢c(P,s) → (Q,t)" 
  shows  "Γ⊢c( normalizec P,s) → (normalizec Q,t)"
using exec
proof(induct rule:stepc_induct)
  case (Catchc P s Q t c2)
    thus ?case
   *) 
(* qed(auto intro: stepc.intros) *)

subsection ‹Equivalence of comp mod semantics and comp mod nested›

definition catch_cond_nest
where
"catch_cond_nest zs Q xs P s s'' s' Γ n ≡ (zs=(map (lift_catch Q) xs) ∨
             ((fst(((P, s)#xs)!length xs)=Throw ∧
               snd(last ((P, s)#xs)) = Normal s' ∧ s=Normal s''∧
               (∃ys. (n,Γ,(Q, snd(((P, s)#xs)!length xs))#ys) ∈ cptn_mod_nest_call ∧ 
                zs=(map (lift_catch Q) xs)@((Q, snd(((P, s)#xs)!length xs))#ys)))) ∨
                ((fst(((P, s)#xs)!length xs)=Skip ∧ 
                 (∃ys. (n,Γ,(Skip,snd(last ((P, s)#xs)))#ys) ∈ cptn_mod_nest_call ∧                   
                 zs=(map (lift_catch Q) xs)@((Skip,snd(last ((P, s)#xs)))#ys)))))
"

lemma div_catch_nest: assumes cptn_m:"(n,Γ,list) ∈ cptn_mod_nest_call"
shows "(∀s P Q zs. list=(Catch P Q, s)#zs ⟶
       (∃xs s' s''. 
          (n, Γ,(P, s)#xs) ∈ cptn_mod_nest_call  ∧ 
             catch_cond_nest zs Q xs P s s'' s' Γ n))  
            "
unfolding catch_cond_nest_def
using cptn_m
proof (induct rule: cptn_mod_nest_call.induct)
case (CptnModNestOne Γ P s)
  thus ?case using cptn_mod_nest_call.CptnModNestOne by blast 
next
  case (CptnModNestSkip  Γ P s t n xs) 
  from CptnModNestSkip.hyps
  have step: "Γ⊢c (P, s) → (Skip, t)" by auto
  from CptnModNestSkip.hyps
  have noskip: "~(P=Skip)" using stepc_elim_cases(1) by blast  
  have no_catch: "∀p1 p2. ¬(P=Catch p1 p2)" using CptnModNestSkip.hyps(2) redex_not_Catch by auto
  from CptnModNestSkip.hyps
  have in_cptn_mod: "(n,Γ, (Skip, t) # xs) ∈ cptn_mod_nest_call" by auto  
  then show ?case using no_catch by simp
next
  case (CptnModNestThrow  Γ P s t n xs) 
  from CptnModNestThrow.hyps
  have step: "Γ⊢c (P, s) → (Throw, t)" by auto 
  from CptnModNestThrow.hyps
  have in_cptn_mod: "(n,Γ, (Throw, t) # xs) ∈ cptn_mod_nest_call" by auto 
  have no_catch: "∀p1 p2. ¬(P=Catch p1 p2)" using CptnModNestThrow.hyps(2) redex_not_Catch by auto
  then show ?case by auto
next
  case (CptnModNestCondT Γ P0 s ys b P1)
  thus ?case using CptnModOne by blast
next
  case (CptnModNestCondF Γ P0 s ys b P1)
  thus ?case using CptnModOne by blast
next
  case (CptnModNestCatch1 sa P Q zs)
  thus ?case by blast
next
  case (CptnModNestCatch2 n Γ P0 s xs ys zs P1) 
  from CptnModNestCatch2.hyps(3) 
  have last:"fst (((P0, s) # xs) ! length xs) = Skip" 
       by (simp add: last_length) 
  have P0cptn:"(n,Γ, (P0, s) # xs) ∈ cptn_mod_nest_call" by fact          
  then have "zs = map (lift_catch P1) xs @((Skip,snd(last ((P0, s)#xs)))#ys)" by (simp add:CptnModNestCatch2.hyps)
  show ?case
  proof -{
    fix sa P Q zsa    
    assume eq:"(Catch P0 P1, s) # zs = (Catch P Q, sa) # zsa"
    then have "P0 =P ∧ P1 = Q ∧ s=sa ∧ zs=zsa" by auto
    then have "(P0, s) = (P, sa)" by auto 
    have "last ((P0, s) # xs) = ((P, sa) # xs) ! length xs"
      by (simp add: `P0 = P ∧ P1 = Q ∧ s = sa ∧ zs = zsa` last_length)
    then have "zs = (map (lift_catch Q) xs)@((Skip,snd(last ((P0, s)#xs)))#ys)"
      using `P0 = P ∧ P1 = Q ∧ s = sa ∧ zs = zsa` `zs = map (lift_catch P1) xs @ ((Skip,snd(last ((P0, s)#xs)))#ys)` 
      by force    
    then have "(∃xs s' s''. ((n,Γ,(P, s)#xs) ∈ cptn_mod_nest_call  ∧ 
             ((zs=(map (lift_catch Q) xs) ∨
             ((fst(((P, s)#xs)!length xs)=Throw ∧
               snd(last ((P, s)#xs)) = Normal s' ∧  s=Normal s''∧
               (∃ys. (n,Γ,(Q, snd(((P, s)#xs)!length xs))#ys) ∈ cptn_mod_nest_call ∧ 
                zs=(map (lift_catch Q) xs)@((Q, snd(((P, s)#xs)!length xs))#ys)))) ∨
                (∃ys. ((fst(((P, s)#xs)!length xs)=Skip ∧ (n,Γ,(Skip,snd(last ((P, s)#xs)))#ys) ∈ cptn_mod_nest_call ∧                 
                 zs=(map (lift_catch Q) xs)@((Skip,snd(last ((P0, s)#xs)))#ys))))))))"
    using P0cptn  `P0 = P ∧ P1 = Q ∧ s = sa ∧ zs = zsa`  last  CptnModNestCatch2.hyps(4) by blast 
   } 
   thus ?thesis by auto
  qed
next
  case (CptnModNestCatch3 n Γ P0 s xs s' P1 ys zs)
  from CptnModNestCatch3.hyps(3)  
  have last:"fst (((P0, Normal s) # xs) ! length xs) = Throw" 
       by (simp add: last_length) 
  from CptnModNestCatch3.hyps(4) 
  have lastnormal:"snd (last ((P0, Normal s) # xs)) = Normal s'"
      by (simp add: last_length)
  have P0cptn:"(n,Γ, (P0, Normal s) # xs) ∈ cptn_mod_nest_call" by fact
  from CptnModNestCatch3.hyps(5) 
    have P1cptn:"(n,Γ, (P1, snd (((P0, Normal s) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call"
      by (simp add: last_length)          
  then have "zs = map (lift_catch P1) xs @ (P1, snd (last ((P0, Normal s) # xs))) # ys" 
    by (simp add:CptnModNestCatch3.hyps)
  show ?case
  proof -{
    fix sa P Q zsa    
    assume eq:"(Catch P0 P1, Normal s) # zs = (Catch P Q, Normal sa) # zsa"
    then have "P0 =P ∧ P1 = Q ∧ Normal s= Normal sa ∧ zs=zsa" by auto     
    have "last ((P0, Normal s) # xs) = ((P, Normal sa) # xs) ! length xs"
      by (simp add: `P0 = P ∧ P1 = Q ∧ Normal s = Normal sa ∧ zs = zsa` last_length)
    then have "zsa = map (lift_catch Q) xs @ (Q, snd (((P, Normal sa) # xs) ! length xs)) # ys"
      using `P0 = P ∧ P1 = Q ∧ Normal s = Normal sa ∧ zs = zsa` `zs = map (lift_catch P1) xs @ (P1, snd (last ((P0, Normal s) # xs))) # ys` by force
    then have "(n,Γ, (P, Normal s) # xs) ∈ cptn_mod_nest_call ∧ (fst(((P, Normal s)#xs)!length xs)=Throw ∧
               snd(last ((P, Normal s)#xs)) = Normal s' ∧ 
              (∃ys. (n,Γ,(Q, snd(((P, Normal s)#xs)!length xs))#ys) ∈ cptn_mod_nest_call ∧ 
              zs=(map (lift_catch Q) xs)@((Q, snd(((P, Normal s)#xs)!length xs))#ys)))" 
      using lastnormal P1cptn P0cptn `P0 = P ∧ P1 = Q ∧ Normal s = Normal sa ∧ zs = zsa` last 
       by auto    
    }note this [of P0 P1 s zs] thus ?thesis by blast qed
next
  case (CptnModNestEnv Γ P s t n xs)  
  then have step:"(n, Γ, (P, t) # xs) ∈ cptn_mod_nest_call" by auto  
  have step_e: "Γ⊢c (P, s) →e (P, t)" using CptnModNestEnv by auto
  show ?case     
    proof (cases P)
      case (Catch P1 P2) 
      then have eq_P_Catch:"(P, t) # xs = (LanguageCon.com.Catch P1 P2, t) # xs" by auto      
      then  obtain xsa t' t'' where 
         p1:"(n,Γ, (P1, t) # xsa) ∈ cptn_mod_nest_call" and 
        p2:" (xs = map (lift_catch P2) xsa ∨
            fst (((P1, t) # xsa) ! length xsa) = LanguageCon.com.Throw ∧
            snd (last ((P1, t) # xsa)) = Normal t' ∧
            t = Normal t'' ∧
            (∃ys. (n,Γ, (P2, snd (((P1, t) # xsa) ! length xsa)) # ys) ∈ cptn_mod_nest_call ∧
              xs = map (lift_catch P2) xsa @ (P2, snd (((P1, t) # xsa) ! length xsa)) # ys) ∨
                fst (((P1, t) # xsa) ! length xsa) = LanguageCon.com.Skip ∧
                (∃ys.(n,Γ,(Skip,snd(last ((P1, t)#xsa)))#ys) ∈ cptn_mod_nest_call ∧ 
                xs = map (lift_catch P2) xsa @
                ((LanguageCon.com.Skip, snd (last ((P1, t) # xsa)))#ys)))" 
        using CptnModNestEnv(3) by auto
      have all_step:"(n,Γ, (P1, s)#((P1, t) # xsa)) ∈ cptn_mod_nest_call"
        using p1 Env Env_n cptn_mod.CptnModEnv env_normal_s step_e
      proof -
        have f1: "SmallStepCon.redex P = SmallStepCon.redex P1"
          using local.Catch by auto
        obtain bb :: "('b, 'c) xstate ⇒ 'b" where
          "∀x2. (∃v5. x2 = Normal v5) = (x2 = Normal (bb x2))"
          by moura
        then have "s = t ∨ s = Normal (bb s)"
          by (metis (no_types) env_normal_s step_e)
        then show ?thesis
          using f1 by (metis (no_types) Env Env_n cptn_mod_nest_call.CptnModNestEnv p1)
      qed
      show ?thesis using p2 
      proof  
        assume "xs = map (lift_catch P2) xsa"
        have "(P, t) # xs = map (lift_catch P2) ((P1, t) # xsa)"
          by (simp add: `xs = map (lift_catch P2) xsa` lift_catch_def local.Catch)
        thus ?thesis using all_step eq_P_Catch by fastforce
      next 
        assume 
         "fst (((P1, t) # xsa) ! length xsa) = LanguageCon.com.Throw ∧
          snd (last ((P1, t) # xsa)) = Normal t' ∧
          t = Normal t'' ∧
          (∃ys. (n,Γ, (P2, snd (((P1, t) # xsa) ! length xsa)) # ys) ∈ cptn_mod_nest_call ∧
                xs =
                map (lift_catch P2) xsa @
                (P2, snd (((P1, t) # xsa) ! length xsa)) # ys) ∨
                fst (((P1, t) # xsa) ! length xsa) = LanguageCon.com.Skip ∧
           (∃ys. (n,Γ,(Skip,snd(last ((P1, t)#xsa)))#ys) ∈ cptn_mod_nest_call ∧ 
            xs = map (lift_catch P2) xsa @
            ((LanguageCon.com.Skip, snd (last ((P1, t) # xsa)))#ys))"      
         then show ?thesis 
         proof
           assume 
            a1:"fst (((P1, t) # xsa) ! length xsa) = LanguageCon.com.Throw ∧
             snd (last ((P1, t) # xsa)) = Normal t' ∧
             t = Normal t'' ∧
             (∃ys. (n,Γ, (P2, snd (((P1, t) # xsa) ! length xsa)) # ys) ∈ cptn_mod_nest_call ∧
                xs = map (lift_catch P2) xsa @
                       (P2, snd (((P1, t) # xsa) ! length xsa)) # ys)"
            then obtain ys where p2_exec:"(n,Γ, (P2, snd (((P1, t) # xsa) ! length xsa)) # ys) ∈ cptn_mod_nest_call ∧
                xs = map (lift_catch P2) xsa @
                       (P2, snd (((P1, t) # xsa) ! length xsa)) # ys" 
            by fastforce
            from a1 obtain t1 where t_normal: "t=Normal t1" 
              using env_normal_s'_normal_s by blast
            have f1:"fst (((P1, s)#(P1, t) # xsa) ! length ((P1, t)#xsa)) = LanguageCon.com.Throw"
              using a1 by fastforce
             from a1 have last_normal: "snd (last ((P1, s)#(P1, t) # xsa)) = Normal t'"
               by fastforce 
             then have p2_long_exec: "(n,Γ, (P2, snd (((P1, s)#(P1, t) # xsa) ! length ((P1, s)#xsa))) # ys) ∈ cptn_mod_nest_call ∧
                (P, t)#xs = map (lift_catch P2) ((P1, t) # xsa) @
                       (P2, snd (((P1, s)#(P1, t) # xsa) ! length ((P1, s)#xsa))) # ys" using p2_exec 
                by (simp add: lift_catch_def local.Catch)                  
             thus ?thesis using  a1 f1 last_normal all_step eq_P_Catch 
               by (clarify, metis (no_types) list.size(4) not_step_c_env  step_e)            
           next
           assume 
            as1:"fst (((P1, t) # xsa) ! length xsa) = LanguageCon.com.Skip ∧
           (∃ys. (n,Γ,(Skip,snd(last ((P1, t)#xsa)))#ys) ∈ cptn_mod_nest_call ∧ 
            xs = map (lift_catch P2) xsa @
            ((LanguageCon.com.Skip, snd (last ((P1, t) # xsa)))#ys))"               
            then obtain ys where p1:"(n,Γ,(Skip,snd(last ((P1, t)#xsa)))#ys) ∈ cptn_mod_nest_call ∧ 
                         (P, t)#xs = map (lift_catch P2) ((P1, t) # xsa) @
                          ((LanguageCon.com.Skip, snd (last ((P1, t) # xsa)))#ys)"
            proof -
              assume a1: "⋀ys. (n,Γ, (LanguageCon.com.Skip, snd (last ((P1, t) # xsa))) # ys) ∈ cptn_mod_nest_call ∧ 
                         (P, t) # xs = map (lift_catch P2) ((P1, t) # xsa) @ 
                         (LanguageCon.com.Skip, snd (last ((P1, t) # xsa))) # ys ⟹ 
                         thesis"
              have "(LanguageCon.com.Catch P1 P2, t) # map (lift_catch P2) xsa = map (lift_catch P2) ((P1, t) # xsa)"
                by (simp add: lift_catch_def)
              thus ?thesis
                using a1 as1 eq_P_Catch by moura
            qed            
            from as1 have p2: "fst (((P1, s)#(P1, t) # xsa) ! length ((P1, t) #xsa)) = LanguageCon.com.Skip"
                 by fastforce                              
            thus ?thesis using p1 all_step eq_P_Catch by fastforce
          qed
      qed
    qed (auto)
qed(force+)


definition seq_cond_nest
where
"seq_cond_nest zs Q xs P s s'' s' Γ n ≡ (zs=(map (lift Q) xs) ∨
             ((fst(((P, s)#xs)!length xs)=Skip ∧ 
               (∃ys. (n,Γ,(Q, snd(((P, s)#xs)!length xs))#ys) ∈ cptn_mod_nest_call ∧ 
                zs=(map (lift (Q)) xs)@((Q, snd(((P, s)#xs)!length xs))#ys)))) ∨
             ((fst(((P, s)#xs)!length xs)=Throw ∧ 
                 snd(last ((P, s)#xs)) = Normal s' ∧  s=Normal s''∧
                 (∃ys.  (n,Γ,(Throw,Normal s')#ys) ∈ cptn_mod_nest_call ∧
                      zs=(map (lift Q) xs)@((Throw,Normal s')#ys)))))
"

lemma div_seq_nest: assumes cptn_m:"(n,Γ,list) ∈ cptn_mod_nest_call"
shows "(∀s P Q zs. list=(Seq P Q, s)#zs ⟶
       (∃xs s' s''. 
          (n,Γ,(P, s)#xs) ∈ cptn_mod_nest_call  ∧ 
             seq_cond_nest zs Q xs P s s'' s' Γ n))  
            "
unfolding seq_cond_nest_def
using cptn_m
proof (induct rule: cptn_mod_nest_call.induct)
  case (CptnModNestOne Γ P s)
  thus ?case using cptn_mod_nest_call.CptnModNestOne 
   by blast
next
  case (CptnModNestSkip  Γ P s t n xs) 
  from CptnModNestSkip.hyps
  have step: "Γ⊢c (P, s) → (Skip, t)" by auto
  from CptnModNestSkip.hyps
  have noskip: "~(P=Skip)" using stepc_elim_cases(1) by blast  
  have x: "∀c c1 c2. redex c = Seq c1 c2 ⟹ False"
          using redex_not_Seq by blast
  from CptnModNestSkip.hyps
  have in_cptn_mod: "(n,Γ, (Skip, t) # xs) ∈ cptn_mod_nest_call" by auto  
  then show ?case using CptnModNestSkip.hyps(2) SmallStepCon.redex_not_Seq by blast
next
  case (CptnModNestThrow  Γ P s t xs)
  from CptnModNestThrow.hyps
  have step: "Γ⊢c (P, s) → (Throw, t)" by auto 
  moreover from CptnModNestThrow.hyps 
  have no_seq: "∀p1 p2. ¬(P=Seq p1 p2)" using CptnModNestThrow.hyps(2) redex_not_Seq by auto
  ultimately show ?case by auto   
next 
  case (CptnModNestCondT Γ P0 s ys b P1)
  thus ?case by auto
next
  case (CptnModNestCondF Γ P0 s ys b P1)
  thus ?case by auto
next
  case (CptnModNestSeq1 n Γ P0 s xs zs P1) thus ?case 
    by blast   
next 
  case (CptnModNestSeq2 n Γ P0 s xs P1 ys zs) 
  from CptnModNestSeq2.hyps(3) last_length have last:"fst (((P0, s) # xs) ! length xs) = Skip" 
       by (simp add: last_length) 
  have P0cptn:"(n,Γ, (P0, s) # xs) ∈ cptn_mod_nest_call" by fact
  from CptnModNestSeq2.hyps(4) have P1cptn:"(n,Γ, (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call"
      by (simp add: last_length)          
  then have "zs = map (lift P1) xs @ (P1, snd (last ((P0, s) # xs))) # ys" by (simp add:CptnModNestSeq2.hyps)
  show ?case
  proof -{
    fix sa P Q zsa    
    assume eq:"(Seq P0 P1, s) # zs = (Seq P Q, sa) # zsa"
    then have "P0 =P ∧ P1 = Q ∧ s=sa ∧ zs=zsa" by auto 
     have "last ((P0, s) # xs) = ((P, sa) # xs) ! length xs"
            by (simp add: `P0 = P ∧ P1 = Q ∧ s = sa ∧ zs = zsa` last_length)
    then have "zsa = map (lift Q) xs @ (Q, snd (((P, sa) # xs) ! length xs)) # ys"
         using `P0 = P ∧ P1 = Q ∧ s = sa ∧ zs = zsa` `zs = map (lift P1) xs @ (P1, snd (last ((P0, s) # xs))) # ys` 
         by force    
    then have "(∃xs s' s''. (n,Γ, (P, sa) # xs) ∈ cptn_mod_nest_call ∧
                        (zsa = map (lift Q) xs ∨              
                         fst (((P, sa) # xs) ! length xs) = Skip ∧
                             (∃ys. (n,Γ, (Q, snd (((P, sa) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
                             zsa = map (lift Q) xs @ (Q, snd (((P, sa) # xs) ! length xs)) # ys) ∨
                        ((fst(((P, sa)#xs)!length xs)=Throw ∧ 
                          snd(last ((P, sa)#xs)) = Normal s' ∧  s=Normal s''∧
                         (∃ys. (n,Γ,(Throw,Normal s')#ys) ∈ cptn_mod_nest_call ∧
                               zsa=(map (lift Q) xs)@((Throw,Normal s')#ys))))))
               " 
        using P0cptn P1cptn  `P0 = P ∧ P1 = Q ∧ s = sa ∧ zs = zsa` last 
        by blast                 
   }  
   thus ?case by auto qed     
next
  case (CptnModNestSeq3 n Γ P0 s xs s' ys zs P1) 
  from CptnModNestSeq3.hyps(3) 
  have last:"fst (((P0, Normal s) # xs) ! length xs) = Throw" 
       by (simp add: last_length) 
  have P0cptn:"(n,Γ, (P0, Normal s) # xs) ∈ cptn_mod_nest_call" by fact
  from CptnModNestSeq3.hyps(4) 
  have lastnormal:"snd (last ((P0, Normal s) # xs)) = Normal s'"
      by (simp add: last_length)          
  then have "zs = map (lift P1) xs @ ((Throw, Normal s')#ys)" by (simp add:CptnModNestSeq3.hyps)
  show ?case
  proof -{
    fix sa P Q zsa    
    assume eq:"(Seq P0 P1, Normal s) # zs = (Seq P Q, Normal sa) # zsa"
    then have "P0 =P ∧ P1 = Q ∧ Normal s=Normal sa ∧ zs=zsa" by auto
    then have "(P0, Normal s) = (P, Normal sa)" by auto 
    have "last ((P0, Normal s) # xs) = ((P, Normal sa) # xs) ! length xs"
                    by (simp add: `P0 = P ∧ P1 = Q ∧ Normal s = Normal sa ∧ zs = zsa` last_length)
    then have zsa:"zsa = (map (lift Q) xs)@((Throw,Normal s')#ys)"
                    using `P0 = P ∧ P1 = Q ∧ Normal s = Normal sa ∧ zs = zsa` `zs = map (lift P1) xs @ ((Throw, Normal s')#ys)` 
    by force
    then have a1:"(n,Γ,(Throw,Normal s')#ys) ∈ cptn_mod_nest_call" using CptnModNestSeq3.hyps(5) by blast               
     have "(P, Normal sa::('b, 'c) xstate) = (P0, Normal s)"
    using `P0 = P ∧ P1 = Q ∧ Normal s = Normal sa ∧ zs = zsa` by auto  
    then have "(∃xs s'. (n,Γ, (P, Normal sa) # xs) ∈ cptn_mod_nest_call ∧
                        (zsa = map (lift Q) xs ∨              
                         fst (((P,Normal sa) # xs) ! length xs) = Skip ∧
                             (∃ys. (n,Γ, (Q, snd (((P, Normal sa) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
                             zsa = map (lift Q) xs @ (Q, snd (((P, Normal sa) # xs) ! length xs)) # ys) ∨
                        ((fst(((P, Normal sa)#xs)!length xs)=Throw ∧ 
                          snd(last ((P, Normal sa)#xs)) = Normal s' ∧
                          (∃ys. (n,Γ,(Throw,Normal s')#ys) ∈ cptn_mod_nest_call ∧
                          zsa=(map (lift Q) xs)@((Throw,Normal s')#ys))))))"
     using P0cptn zsa a1 last lastnormal 
       by blast 
   }  
   thus ?thesis by auto qed  
next 
  case (CptnModNestEnv  Γ P s t n zs) 
  then have step:"(n,Γ, (P, t) # zs) ∈ cptn_mod_nest_call" by auto
  have step_e: "Γ⊢c (P, s) →e (P, t)" using CptnModNestEnv by auto  
  show ?case     
    proof (cases P)
      case (Seq P1 P2) 
      then have eq_P:"(P, t) # zs = (LanguageCon.com.Seq P1 P2, t) # zs" by auto      
      then  obtain xs t' t'' where 
         p1:"(n,Γ, (P1, t) # xs) ∈ cptn_mod_nest_call" and p2:"
     (zs = map (lift P2) xs ∨
      fst (((P1, t) # xs) ! length xs) = LanguageCon.com.Skip ∧      
      (∃ys. (n,Γ, (P2, snd (((P1, t) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
            zs =
            map (lift P2) xs @
            (P2, snd (((P1, t) # xs) ! length xs)) # ys) ∨
      fst (((P1, t) # xs) ! length xs) = LanguageCon.com.Throw ∧
      snd (last ((P1, t) # xs)) = Normal t' ∧
      t = Normal t'' ∧ (∃ys. (n,Γ,(Throw,Normal t')#ys) ∈ cptn_mod_nest_call ∧ 
      zs =
      map (lift P2) xs @
      ((LanguageCon.com.Throw, Normal t')#ys))) " 
        using CptnModNestEnv(3) by auto
      have all_step:"(n,Γ, (P1, s)#((P1, t) # xs)) ∈ cptn_mod_nest_call" 
        using p1 Env Env_n cptn_mod_nest_call.CptnModNestEnv env_normal_s step_e
      proof -
        have "SmallStepCon.redex P = SmallStepCon.redex P1"
          by (metis SmallStepCon.redex.simps(4) local.Seq)
        then show ?thesis
          by (metis (no_types) Env Env_n cptn_mod_nest_call.CptnModNestEnv env_normal_s p1 step_e)
      qed             
      show ?thesis using p2 
      proof  
        assume "zs = map (lift P2) xs"
        have "(P, t) # zs = map (lift P2) ((P1, t) # xs)"
          by (simp add: `zs = map (lift P2) xs` lift_def local.Seq)
        thus ?thesis using all_step eq_P by fastforce
      next 
        assume 
         "fst (((P1, t) # xs) ! length xs) = LanguageCon.com.Skip ∧      
         (∃ys. (n,Γ, (P2, snd (((P1, t) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
            zs = map (lift P2) xs @ (P2, snd (((P1, t) # xs) ! length xs)) # ys) ∨
          fst (((P1, t) # xs) ! length xs) = LanguageCon.com.Throw ∧
           snd (last ((P1, t) # xs)) = Normal t' ∧
           t = Normal t'' ∧ (∃ys. (n,Γ,(Throw,Normal t')#ys) ∈ cptn_mod_nest_call ∧ 
           zs = map (lift P2) xs @ ((LanguageCon.com.Throw, Normal t')#ys))"
         then show ?thesis 
         proof
           assume 
            a1:"fst (((P1, t) # xs) ! length xs) = LanguageCon.com.Skip ∧      
               (∃ys. (n,Γ, (P2, snd (((P1, t) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
               zs = map (lift P2) xs @ (P2, snd (((P1, t) # xs) ! length xs)) # ys)"
               from a1 obtain ys where 
                   p2_exec:"(n,Γ, (P2, snd (((P1, t) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
                       zs = map (lift P2) xs @
                       (P2, snd (((P1, t) # xs) ! length xs)) # ys" 
                by auto 
               have f1:"fst (((P1, s)#(P1, t) # xs) ! length ((P1, t)#xs)) = LanguageCon.com.Skip"
                 using a1 by fastforce             
               then have p2_long_exec: 
                 "(n,Γ, (P2, snd (((P1, s)#(P1, t) # xs) ! length ((P1, t)#xs))) # ys) ∈ cptn_mod_nest_call ∧
                  (P, t)#zs = map (lift P2) ((P1, t) # xs) @
                       (P2, snd (((P1, s)#(P1, t) # xs) ! length ((P1, t)#xs))) # ys" 
             using p2_exec by (simp add: lift_def local.Seq)     
             thus ?thesis using a1 f1 all_step eq_P by blast            
           next
           assume 
            a1:"fst (((P1, t) # xs) ! length xs) = LanguageCon.com.Throw ∧
            snd (last ((P1, t) # xs)) = Normal t' ∧ t = Normal t'' ∧ 
          (∃ys. (n,Γ,(Throw,Normal t')#ys) ∈ cptn_mod_nest_call ∧ 
           zs = map (lift P2) xs @ ((LanguageCon.com.Throw, Normal t')#ys))"             
             then have last_throw:
               "fst (((P1, s)#(P1, t) # xs) ! length ((P1, t) #xs)) = LanguageCon.com.Throw" 
               by fastforce
             from a1 have last_normal: "snd (last ((P1, s)#(P1, t) # xs)) = Normal t'"
               by fastforce
             have seq_lift:
               "(LanguageCon.com.Seq P1 P2, t) # map (lift P2) xs = map (lift P2) ((P1, t) # xs)"
                by (simp add: a1 lift_def)             
            thus  ?thesis using a1 last_throw last_normal all_step eq_P         
              by (clarify, metis (no_types, lifting) append_Cons env_normal_s'_normal_s  step_e)                 
          qed
      qed
    qed (auto) 
qed (force)+

lemma map_lift_eq_xs_xs':"map (lift a) xs = map (lift a) xs' ⟹ xs=xs'" 
proof (induct xs arbitrary: xs')
  case Nil thus ?case by auto
next
  case (Cons x xsa)
  then have a0:"(lift a) x # map (lift a) xsa = map (lift a) (x # xsa)"
    by fastforce 
  also obtain x' xsa' where xs':"xs' = x'#xsa'" 
    using Cons by auto
  ultimately have a1:"map (lift a) (x # xsa) =map (lift a) (x' # xsa')"
    using Cons by auto
  then have xs:"xsa=xsa'" using a0 a1 Cons by fastforce
  then have "(lift a) x' = (lift a) x" using a0 a1  by auto
  then have "x' = x" unfolding lift_def
    by (metis (no_types, lifting) LanguageCon.com.inject(3) 
               case_prod_beta old.prod.inject prod.collapse)   
  thus ?case using xs xs' by auto
qed

lemma map_lift_catch_eq_xs_xs':"map (lift_catch a) xs = map (lift_catch a) xs' ⟹ xs=xs'" 
proof (induct xs arbitrary: xs')
  case Nil thus ?case by auto
next
  case (Cons x xsa)
  then have a0:"(lift_catch a) x # map (lift_catch a) xsa = map (lift_catch a) (x # xsa)"
    by auto 
  also obtain x' xsa' where xs':"xs' = x'#xsa'" 
    using Cons by auto
  ultimately have a1:"map (lift_catch a) (x # xsa) =map (lift_catch a) (x' # xsa')"
    using Cons by auto  
  then have xs:"xsa=xsa'" using a0 a1 Cons by fastforce  
  then have "(lift_catch a) x' = (lift_catch a) x" using a0 a1  by auto
  then have "x' = x" unfolding lift_catch_def
    by (metis (no_types, lifting) LanguageCon.com.inject(9) 
               case_prod_beta old.prod.inject prod.collapse)   
  thus ?case using xs xs' by auto
qed

lemma map_lift_all_seq:
 assumes a0:"zs=map (lift a) xs" and
         a1:"i<length zs"
 shows "∃b. fst (zs!i) = Seq b a"
using a0 a1
proof (induct zs arbitrary: xs i)
  case Nil thus ?case by auto
next
  case (Cons z1 zsa) thus ?case unfolding lift_def
  proof -
    assume a1: "z1 # zsa = map (λb. case b of (P, s) ⇒ (LanguageCon.com.Seq P a, s)) xs"
    have "∀p c. ∃x. ∀pa ca xa. 
            (pa ≠ (ca::('a, 'b, 'c, 'd) LanguageCon.com, xa::('a, 'c) xstate) ∨ ca = fst pa) ∧ 
            ((c::('a, 'b, 'c, 'd) LanguageCon.com) ≠ fst p ∨ (c, x::('a, 'c) xstate) = p)"
      by fastforce
    then obtain xx :: "('a, 'b, 'c, 'd) LanguageCon.com × ('a, 'c) xstate ⇒ ('a, 'b, 'c, 'd) LanguageCon.com ⇒ ('a, 'c) xstate" where
      "⋀p c x ca pa. (p ≠ (c::('a, 'b, 'c, 'd) LanguageCon.com, x::('a, 'c) xstate) ∨ c = fst p) ∧ (ca ≠ fst pa ∨ (ca, xx pa ca) = pa)"
      by (metis (full_types))  
    then show ?thesis
      using a1 ‹i < length (z1 # zsa)›
      by (simp add: Cons.hyps Cons.prems(1) case_prod_beta')
  qed
qed

lemma map_lift_catch_all_catch:
 assumes a0:"zs=map (lift_catch a) xs" and
         a1:"i<length zs"
 shows "∃b. fst (zs!i) = Catch b a"
using a0 a1
proof (induct zs arbitrary: xs i)
  case Nil thus ?case by auto
next
  case (Cons z1 zsa) thus ?case unfolding lift_catch_def
  proof -
    assume a1: "z1 # zsa = map (λb. case b of (P, s) ⇒ (LanguageCon.com.Catch P a, s)) xs"
    have "∀p c. ∃x. ∀pa ca xa. 
            (pa ≠ (ca::('a, 'b, 'c, 'd) LanguageCon.com, xa::('a, 'c) xstate) ∨ ca = fst pa) ∧ 
            ((c::('a, 'b, 'c, 'd) LanguageCon.com) ≠ fst p ∨ (c, x::('a, 'c) xstate) = p)"
      by fastforce
    then obtain xx :: "('a, 'b, 'c, 'd) LanguageCon.com × ('a, 'c) xstate ⇒ ('a, 'b, 'c, 'd) LanguageCon.com ⇒ ('a, 'c) xstate" where
      "⋀p c x ca pa. (p ≠ (c::('a, 'b, 'c, 'd) LanguageCon.com, x::('a, 'c) xstate) ∨ c = fst p) ∧ (ca ≠ fst pa ∨ (ca, xx pa ca) = pa)"
      by (metis (full_types))  
    then show ?thesis
      using a1 ‹i < length (z1 # zsa)›
      by (simp add: Cons.hyps Cons.prems(1) case_prod_beta')
  qed
qed

lemma map_lift_some_eq_pos:
 assumes a0:"map (lift P) xs @ (P1, s1)#ys = 
             map (lift P) xs'@ (P2, s2)#ys'" and
         a1:"∀p0. P1≠Seq p0 P" and
         a2:"∀p0. P2≠Seq p0 P" 
 shows "length xs = length xs'"
proof -
  {assume ass:"length xs ≠ length xs'"
   { assume ass:"length xs < length xs'"
     then have False using a0  map_lift_all_seq a1 a2
     by (metis (no_types, lifting) fst_conv length_map nth_append nth_append_length)
   }note l=this
   { assume ass:"length xs > length xs'"
     then have False using a0  map_lift_all_seq a1 a2
     by (metis (no_types, lifting) fst_conv length_map nth_append nth_append_length)
   }  then have False using l ass by fastforce
  }
  thus ?thesis by auto
qed

lemma map_lift_some_eq:
 assumes a0:"map (lift P) xs @ (P1, s1)#ys = 
             map (lift P) xs'@ (P2, s2)#ys'" and
        a1:"∀p0. P1≠Seq p0 P" and
        a2:"∀p0. P2≠Seq p0 P" 
 shows "xs' = xs ∧ ys = ys'"
proof -
  have "length xs = length xs'" using a0 map_lift_some_eq_pos a1 a2 by blast
  also have "xs' = xs" using a0 assms calculation map_lift_eq_xs_xs' by fastforce
  ultimately show ?thesis using a0 by fastforce
qed

lemma map_lift_catch_some_eq_pos:
 assumes a0:"map (lift_catch P) xs @ (P1, s1)#ys = 
             map (lift_catch P) xs'@ (P2, s2)#ys'" and
         a1:"∀p0. P1≠Catch p0 P" and
         a2:"∀p0. P2≠Catch p0 P" 
 shows "length xs = length xs'"
proof -
  {assume ass:"length xs ≠ length xs'"
   { assume ass:"length xs < length xs'"
     then have False using a0  map_lift_catch_all_catch a1 a2
     by (metis (no_types, lifting) fst_conv length_map nth_append nth_append_length)
   }note l=this
   { assume ass:"length xs > length xs'"
     then have False using a0  map_lift_catch_all_catch a1 a2
     by (metis (no_types, lifting) fst_conv length_map nth_append nth_append_length)
   }  then have False using l ass by fastforce
  }
  thus ?thesis by auto
qed

lemma map_lift_catch_some_eq:
 assumes a0:"map (lift_catch P) xs @ (P1, s1)#ys = 
             map (lift_catch P) xs'@ (P2, s2)#ys'" and
        a1:"∀p0. P1≠Catch p0 P" and
        a2:"∀p0. P2≠Catch p0 P" 
 shows "xs' = xs ∧ ys = ys'"
proof -
  have "length xs = length xs'" using a0 map_lift_catch_some_eq_pos a1 a2 by blast
  also have "xs' = xs" using a0 assms calculation map_lift_catch_eq_xs_xs' by fastforce
  ultimately show ?thesis using a0 by fastforce
qed


lemma Seq_P_Not_finish:
 assumes
   a0:"zs = map (lift Q) xs" and
   a1:"(m, Γ,(LanguageCon.com.Seq P Q, s) # zs) ∈ cptn_mod_nest_call" and
   a2:"seq_cond_nest zs Q xs' P s s'' s' Γ m" 
shows "xs=xs'"
using a2 unfolding seq_cond_nest_def 
proof
  assume "zs= map (lift Q) xs'"
  then have  "map (lift Q) xs' = 
              map (lift Q) xs" using a0 by auto 
  thus ?thesis using map_lift_eq_xs_xs' by fastforce
next
  assume 
    ass:"fst (((P, s) # xs') ! length xs') = LanguageCon.com.Skip ∧
         (∃ys. (m, Γ, (Q, snd (((P, s) # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift Q) xs' @ (Q, snd (((P, s) # xs') ! length xs')) # ys) ∨
         fst (((P, s) # xs') ! length xs') = LanguageCon.com.Throw ∧
         snd (last ((P, s) # xs')) = Normal s' ∧
         s = Normal s'' ∧
         (∃ys. (m, Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift Q) xs' @ (LanguageCon.com.Throw, Normal s') # ys)"
   {assume 
     ass:"fst (((P, s) # xs') ! length xs') = LanguageCon.com.Skip ∧
        (∃ys. (m, Γ, (Q, snd (((P, s) # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift Q) xs' @ (Q, snd (((P, s) # xs') ! length xs')) # ys)"
      then obtain ys where 
         zs:"zs = map (lift Q) xs' @ (Q, snd (((P, s) # xs') ! length xs')) # ys" 
             by auto 
      then have zs_while:"fst (zs!(length (map (lift Q) xs'))) =
                   Q"  by (metis fstI nth_append_length) 
      have "length zs = length (map (lift Q) xs' @
         (Q, snd (((P, s) # xs') ! length xs')) # ys)" 
          using zs by auto
      then have "(length (map (lift Q) xs')) < 
                  length zs" by auto       
      then have ?thesis using a0 zs_while map_lift_all_seq
         using seq_and_if_not_eq(4) by fastforce     
   }note l = this
   {assume ass:"fst (((P, s) # xs') ! length xs') = LanguageCon.com.Throw ∧
         snd (last ((P, s) # xs')) = Normal s' ∧
         s = Normal s'' ∧
         (∃ys. (m, Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift Q) xs' @ (LanguageCon.com.Throw, Normal s') # ys)"
      then obtain ys where 
            zs:"zs = map (lift Q) xs' @ 
                 (LanguageCon.com.Throw, Normal s') # ys" by auto
      then have zs_while:
          "fst (zs!(length (map (lift Q) xs'))) = Throw"  by (metis fstI nth_append_length) 
       have "length zs = length (map (lift Q) xs' @(LanguageCon.com.Throw, Normal s') # ys)" 
           using zs by auto
       then have "(length (map (lift Q) xs')) < 
                  length zs" by auto       
       then have ?thesis using a0 zs_while map_lift_all_seq
         using seq_and_if_not_eq(4) by fastforce
   } thus ?thesis using l ass by auto
qed

lemma Seq_P_Ends_Normal:
 assumes
   a0:"zs = map (lift Q) xs @ (Q, snd (last ((P, s) # xs))) # ys" and
   a0':"fst (last ((P, s) # xs)) = Skip" and
   a1:"(m, Γ,(LanguageCon.com.Seq P Q, s) # zs) ∈ cptn_mod_nest_call" and
   a2:"seq_cond_nest zs Q xs' P s s'' s' Γ m" 
shows "xs=xs' ∧ (m,Γ,(Q, snd(((P, s)#xs)!length xs))#ys) ∈ cptn_mod_nest_call"
using a2 unfolding seq_cond_nest_def 
proof
  assume ass:"zs= map (lift Q) xs'"
  then have  "map (lift Q) xs' = 
              map (lift Q) xs @ (Q, snd (last ((P, s) # xs))) # ys" using a0 by auto 
  then have zs_while:"fst (zs!(length (map (lift Q) xs))) = Q"  
    by (metis a0 fstI nth_append_length) 
  also have "length zs = 
             length (map (lift Q) xs @ (Q, snd (last ((P, s) # xs))) # ys)" 
    using a0 by auto
  then have "(length (map (lift Q) xs)) < length zs" by auto       
  then show ?thesis using ass zs_while map_lift_all_seq
           using seq_and_if_not_eq(4)
  by metis   
next
  assume 
    ass:"fst (((P, s) # xs') ! length xs') = LanguageCon.com.Skip ∧
         (∃ys. (m, Γ, (Q, snd (((P, s) # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift Q) xs' @ (Q, snd (((P, s) # xs') ! length xs')) # ys) ∨
         fst (((P, s) # xs') ! length xs') = LanguageCon.com.Throw ∧
         snd (last ((P, s) # xs')) = Normal s' ∧
         s = Normal s'' ∧
         (∃ys. (m, Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift Q) xs' @ (LanguageCon.com.Throw, Normal s') # ys)"
   {assume 
     ass:"fst (((P, s) # xs') ! length xs') = LanguageCon.com.Skip ∧
        (∃ys. (m, Γ, (Q, snd (((P, s) # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift Q) xs' @ (Q, snd (((P, s) # xs') ! length xs')) # ys)"
      then obtain ys' where 
         zs:"zs = map (lift Q) xs' @ (Q, snd (((P, s) # xs') ! length xs')) # ys' ∧
             (m, Γ, (Q, snd (((P, s) # xs') ! length xs')) # ys') ∈ cptn_mod_nest_call" 
             by auto 
      then have ?thesis
        using  map_lift_some_eq[of Q xs Q _ ys xs' Q _ ys'] 
               zs  a0 seq_and_if_not_eq(4)[of Q] 
        by auto               
   }note l = this
   {assume ass:"fst (((P, s) # xs') ! length xs') = LanguageCon.com.Throw ∧
         snd (last ((P, s) # xs')) = Normal s' ∧
         s = Normal s'' ∧
         (∃ys. (m, Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift Q) xs' @ (LanguageCon.com.Throw, Normal s') # ys)"
      then obtain ys' where 
         zs:"zs = map (lift Q) xs' @ (LanguageCon.com.Throw, Normal s') # ys' ∧
             (m, Γ, (LanguageCon.com.Throw, Normal s') # ys') ∈ cptn_mod_nest_call" 
         by auto
      then have zs_while:
          "fst (zs!(length (map (lift Q) xs'))) = Throw"  by (metis fstI nth_append_length)       
      have False
        by (metis (no_types) LanguageCon.com.distinct(17) 
              LanguageCon.com.distinct(71) 
              a0 a0' ass last_length
              map_lift_some_eq seq_and_if_not_eq(4) zs)
      then have ?thesis
        by metis
   } thus ?thesis using l ass by auto
qed

lemma Seq_P_Ends_Abort:
 assumes
   a0:"zs = map (lift Q) xs @ (Throw, Normal s') # ys" and
   a0':"fst (last ((P, Normal s) # xs)) = Throw" and
   a0'':"snd(last ((P, Normal s) # xs)) = Normal s'" and
   a1:"(m, Γ,(LanguageCon.com.Seq P Q, Normal s) # zs) ∈ cptn_mod_nest_call" and
   a2:"seq_cond_nest zs Q xs' P (Normal s) ns'' ns' Γ m" 
shows "xs=xs' ∧ (m,Γ,(Throw,Normal s')#ys) ∈ cptn_mod_nest_call"
using a2 unfolding seq_cond_nest_def 
proof
  assume ass:"zs= map (lift Q) xs'"
  then have  "map (lift Q) xs' = 
              map (lift Q) xs @ (Throw, Normal s') # ys" using a0 by auto 
  then have zs_while:"fst (zs!(length (map (lift Q) xs))) = Throw"  
    by (metis a0  fstI nth_append_length) 
  also have "length zs = 
             length (map (lift Q) xs @ (Throw, Normal s') # ys)" 
    using a0 by auto
  then have "(length (map (lift Q) xs)) < length zs" by auto       
  then show ?thesis using ass zs_while map_lift_all_seq    
    by (metis (no_types) LanguageCon.com.simps(82))    
next
  assume 
    ass:"fst (((P, Normal s) # xs') ! length xs') = LanguageCon.com.Skip ∧
         (∃ys. (m, Γ, (Q, snd (((P, Normal s) # xs') ! length xs')) # ys)
          ∈ cptn_mod_nest_call ∧
         zs = map (lift Q) xs' @
              (Q, snd (((P, Normal s) # xs') ! length xs')) # ys) ∨
         fst (((P, Normal s) # xs') ! length xs') = LanguageCon.com.Throw ∧
         snd (last ((P, Normal s) # xs')) = Normal ns' ∧
         Normal s = Normal ns'' ∧
         (∃ys. (m, Γ, (LanguageCon.com.Throw, Normal ns') # ys) ∈ cptn_mod_nest_call ∧
            zs = map (lift Q) xs' @ (LanguageCon.com.Throw, Normal ns') # ys)"
   {assume 
     ass:"fst (((P, Normal s) # xs') ! length xs') = LanguageCon.com.Skip ∧
         (∃ys. (m, Γ, (Q, snd (((P, Normal s) # xs') ! length xs')) # ys)
          ∈ cptn_mod_nest_call ∧
         zs = map (lift Q) xs' @
              (Q, snd (((P, Normal s) # xs') ! length xs')) # ys)"
      then obtain ys' where 
         zs:"(m, Γ, (Q, snd (((P, Normal s) # xs') ! length xs')) # ys')
               ∈ cptn_mod_nest_call ∧
            zs = map (lift Q) xs' @ 
              (Q, snd (((P, Normal s) # xs') ! length xs')) # ys'" 
             by auto 
      then have ?thesis
        using a0 seq_and_if_not_eq(4)[of Q]
         by (metis LanguageCon.com.distinct(17) LanguageCon.com.distinct(71) 
             a0' ass last_length map_lift_some_eq)                        
   }note l = this
   {assume ass:"fst (((P, Normal s) # xs') ! length xs') = LanguageCon.com.Throw ∧
         snd (last ((P, Normal s) # xs')) = Normal ns' ∧
         Normal s = Normal ns'' ∧
         (∃ys. (m, Γ, (LanguageCon.com.Throw, Normal ns') # ys) ∈ cptn_mod_nest_call ∧
            zs = map (lift Q) xs' @ (LanguageCon.com.Throw, Normal ns') # ys)"
      then obtain ys' where 
         zs:"(m, Γ, (LanguageCon.com.Throw, Normal ns') # ys') ∈ cptn_mod_nest_call ∧
            zs = map (lift Q) xs' @ (LanguageCon.com.Throw, Normal ns') # ys'" 
         by auto
      then have zs_while:
          "fst (zs!(length (map (lift Q) xs'))) = Throw"  
        by (metis fstI nth_append_length)             
      then have ?thesis using a0 ass map_lift_some_eq by blast        
   } thus ?thesis using l ass by auto
qed

lemma Catch_P_Not_finish:
 assumes
   a0:"zs = map (lift_catch Q) xs" and   
   a1:"catch_cond_nest zs Q xs' P s s'' s' Γ m" 
shows "xs=xs'"
using a1 unfolding catch_cond_nest_def 
proof
  assume "zs= map (lift_catch Q) xs'"
  then have  "map (lift_catch Q) xs' = 
              map (lift_catch Q) xs" using a0 by auto 
  thus ?thesis using map_lift_catch_eq_xs_xs' by fastforce
next
  assume 
    ass:"
         fst (((P, s) # xs') ! length xs') = LanguageCon.com.Throw ∧
         snd (last ((P, s) # xs')) = Normal s' ∧
         s = Normal s'' ∧
         (∃ys. (m, Γ, (Q, snd (((P, s) # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift_catch Q) xs' @ (Q, snd (((P, s) # xs') ! length xs')) # ys) ∨
         fst (((P, s) # xs') ! length xs') = LanguageCon.com.Skip ∧
         (∃ys. (m, Γ, (LanguageCon.com.Skip, snd (last ((P, s) # xs'))) # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift_catch Q) xs' @ (LanguageCon.com.Skip, snd (last ((P, s) # xs'))) # ys)"
   {assume 
     ass:"fst (((P, s) # xs') ! length xs') = LanguageCon.com.Skip ∧
         (∃ys. (m, Γ, (LanguageCon.com.Skip, snd (last ((P, s) # xs'))) # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift_catch Q) xs' @ (LanguageCon.com.Skip, snd (last ((P, s) # xs'))) # ys)"
      then obtain ys where 
         zs:"(m, Γ, (LanguageCon.com.Skip, snd (last ((P, s) # xs'))) # ys) ∈ cptn_mod_nest_call ∧
            zs = map (lift_catch Q) xs' @ (LanguageCon.com.Skip, snd (last ((P, s) # xs'))) # ys" 
             by auto 
      then have zs_while:"fst (zs!(length (map (lift_catch Q) xs'))) = Skip "  
        by (metis fstI nth_append_length) 
      have "length zs = length (map (lift Q) xs' @
         (Q, snd (((P, s) # xs') ! length xs')) # ys)" 
          using zs by auto
      then have "(length (map (lift Q) xs')) < 
                  length zs" by auto       
      then have ?thesis using a0 zs_while map_lift_catch_all_catch
         using seq_and_if_not_eq(12) by fastforce     
   }note l = this
   {assume ass:"fst (((P, s) # xs') ! length xs') = LanguageCon.com.Throw ∧
         snd (last ((P, s) # xs')) = Normal s' ∧
         s = Normal s'' ∧
         (∃ys. (m, Γ, (Q, snd (((P, s) # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift_catch Q) xs' @ (Q, snd (((P, s) # xs') ! length xs')) # ys)"
      then obtain ys where 
            zs:"zs = map (lift_catch Q) xs' @ (Q, snd (((P, s) # xs') ! length xs')) # ys" by auto
      then have zs_while:
        "fst (zs!(length (map (lift Q) xs'))) = Q"
         by (metis (no_types) eq_fst_iff length_map nth_append_length zs) 
       have "length zs = length (map (lift Q) xs' @(LanguageCon.com.Throw, Normal s') # ys)" 
           using zs by auto
       then have "(length (map (lift Q) xs')) < 
                  length zs" by auto       
       then have ?thesis using a0 zs_while map_lift_catch_all_catch
         by fastforce
   } thus ?thesis using l ass by auto
qed

lemma Catch_P_Ends_Normal:
 assumes
   a0:"zs = map (lift_catch Q) xs @ (Q, snd (last ((P, Normal s) # xs))) # ys" and
   a0':"fst (last ((P, Normal s) # xs)) = Throw" and
   a0'':"snd (last ((P, Normal s) # xs)) = Normal s'" and
   a1:"catch_cond_nest zs Q xs' P (Normal s) ns'' ns' Γ m" 
shows "xs=xs' ∧ (m,Γ,(Q, snd(((P, Normal s)#xs)!length xs))#ys) ∈ cptn_mod_nest_call"
using a1 unfolding catch_cond_nest_def 
proof
  assume ass:"zs= map (lift_catch Q) xs'"
  then have  "map (lift_catch Q) xs' = 
              map (lift_catch Q) xs @ (Q, snd (last ((P, Normal s) # xs))) # ys" using a0 by auto 
  then have zs_while:"fst (zs!(length (map (lift_catch Q) xs))) = Q"
    by (metis a0 fstI nth_append_length)      
  also have "length zs = 
             length (map (lift_catch Q) xs @ (Q, snd (last ((P, Normal s) # xs))) # ys)" 
    using a0 by auto
  then have "(length (map (lift_catch Q) xs)) < length zs" by auto       
  then show ?thesis using ass zs_while map_lift_catch_all_catch
           using seq_and_if_not_eq(12)
  by metis
next
  assume 
    ass:"fst (((P, Normal s) # xs') ! length xs') = LanguageCon.com.Throw ∧
         snd (last ((P, Normal s) # xs')) = Normal ns' ∧
         Normal s = Normal ns'' ∧
         (∃ys. (m, Γ, (Q, snd (((P, Normal s) # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift_catch Q) xs' @ (Q, snd (((P, Normal s) # xs') ! length xs')) # ys) ∨
         fst (((P, Normal s) # xs') ! length xs') = LanguageCon.com.Skip ∧
         (∃ys. (m, Γ, (LanguageCon.com.Skip, snd (last ((P, Normal s) # xs'))) # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift_catch Q) xs' @ (LanguageCon.com.Skip, snd (last ((P, Normal s) # xs'))) # ys)"
   {assume 
     ass:"fst (((P, Normal s) # xs') ! length xs') = LanguageCon.com.Skip ∧
         (∃ys. (m, Γ, (LanguageCon.com.Skip, snd (last ((P, Normal s) # xs'))) # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift_catch Q) xs' @ (LanguageCon.com.Skip, snd (last ((P, Normal s) # xs'))) # ys)"
      then obtain ys' where 
         zs:"(m, Γ, (LanguageCon.com.Skip, snd (last ((P, Normal s) # xs'))) # ys') ∈ cptn_mod_nest_call ∧
             zs = map (lift_catch Q) xs' @ (LanguageCon.com.Skip, snd (last ((P, Normal s) # xs'))) # ys'" 
             by auto 
      then have ?thesis
        using  map_lift_catch_some_eq[of Q xs Q _ ys xs' Skip _ ys'] 
               zs  a0 seq_and_if_not_eq(12)[of Q]
        by (metis LanguageCon.com.distinct(17) LanguageCon.com.distinct(19) a0' ass last_length)                        
   }note l = this
   {assume ass:"fst (((P, Normal s) # xs') ! length xs') = LanguageCon.com.Throw ∧
                snd (last ((P, Normal s) # xs')) = Normal ns' ∧
                Normal s = Normal ns'' ∧
                (∃ys. (m, Γ, (Q, snd (((P, Normal s) # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
                zs = map (lift_catch Q) xs' @ (Q, snd (((P, Normal s) # xs') ! length xs')) # ys)"
      then obtain ys' where 
         zs:"(m, Γ, (Q, snd (((P, Normal s) # xs') ! length xs')) # ys') ∈ cptn_mod_nest_call ∧
                zs = map (lift_catch Q) xs' @ (Q, snd (((P, Normal s) # xs') ! length xs')) # ys'" 
         by auto
      then have zs_while:
          "fst (zs!(length (map (lift_catch Q) xs'))) = Q"  by (metis fstI nth_append_length)       

      then have ?thesis 
        using LanguageCon.com.distinct(17) LanguageCon.com.distinct(71) 
            a0 a0' ass last_length map_lift_catch_some_eq[of Q xs Q _ ys xs' Q _ ys']  
            seq_and_if_not_eq(12) zs
        by blast        
   } thus ?thesis using l ass by auto
qed
 
lemma Catch_P_Ends_Skip:
 assumes
   a0:"zs = map (lift_catch Q) xs @ (Skip, snd (last ((P, s) # xs))) # ys" and
   a0':"fst (last ((P,s) # xs)) = Skip" and
   a1:"catch_cond_nest zs Q xs' P s ns'' ns' Γ m" 
shows "xs=xs' ∧ (m,Γ,(Skip,snd(last ((P,s) # xs)))#ys) ∈ cptn_mod_nest_call"
using a1 unfolding catch_cond_nest_def 
proof
  assume ass:"zs= map (lift_catch Q) xs'"
  then have  "map (lift_catch Q) xs' = 
              map (lift_catch Q) xs @ (Skip, snd (last ((P, s) # xs))) # ys" using a0 by auto 
  then have zs_while:"fst (zs!(length (map (lift_catch Q) xs))) = Skip"  
    by (metis a0  fstI nth_append_length) 
  also have "length zs = 
             length (map (lift_catch Q) xs @ (Skip, snd (last ((P, s) # xs))) # ys)" 
    using a0 by auto
  then have "(length (map (lift_catch Q) xs)) < length zs" by auto       
  then show ?thesis using ass zs_while map_lift_catch_all_catch
    by (metis LanguageCon.com.distinct(19))    
next
  assume 
    ass:"fst (((P, s) # xs') ! length xs') = LanguageCon.com.Throw ∧
         snd (last ((P, s) # xs')) = Normal ns' ∧
         s = Normal ns'' ∧
         (∃ys. (m, Γ, (Q, snd (((P, s) # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift_catch Q) xs' @ (Q, snd (((P, s) # xs') ! length xs')) # ys) ∨
         fst (((P, s) # xs') ! length xs') = LanguageCon.com.Skip ∧
         (∃ys. (m, Γ, (LanguageCon.com.Skip, snd (last ((P, s) # xs'))) # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift_catch Q) xs' @ (LanguageCon.com.Skip, snd (last ((P, s) # xs'))) # ys)"
   {assume 
     ass:"fst (((P, s) # xs') ! length xs') = LanguageCon.com.Skip ∧
         (∃ys. (m, Γ, (LanguageCon.com.Skip, snd (last ((P, s) # xs'))) # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift_catch Q) xs' @ (LanguageCon.com.Skip, snd (last ((P, s) # xs'))) # ys)"
      then obtain ys' where 
         zs:"(m, Γ, (LanguageCon.com.Skip, snd (last ((P, s) # xs'))) # ys') ∈ cptn_mod_nest_call ∧
             zs = map (lift_catch Q) xs' @ (LanguageCon.com.Skip, snd (last ((P, s) # xs'))) # ys'" 
             by auto 
      then have ?thesis
        using a0 seq_and_if_not_eq(12)[of Q] a0' ass last_length map_lift_catch_some_eq
        using LanguageCon.com.distinct(19) by blast                
   }note l = this
   {assume ass:"fst (((P, s) # xs') ! length xs') = LanguageCon.com.Throw ∧
         snd (last ((P, s) # xs')) = Normal ns' ∧
         s = Normal ns'' ∧
         (∃ys. (m, Γ, (Q, snd (((P, s) # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
         zs = map (lift_catch Q) xs' @ (Q, snd (((P, s) # xs') ! length xs')) # ys)"
      then obtain ys' where 
         zs:"(m, Γ, (Q, snd (((P, s) # xs') ! length xs')) # ys') ∈ cptn_mod_nest_call ∧
         zs = map (lift_catch Q) xs' @ (Q, snd (((P, s) # xs') ! length xs')) # ys'" 
         by auto
      then have zs_while:
          "fst (zs!(length (map (lift_catch Q) xs'))) = Q"  
        by (metis fstI nth_append_length)             
      then have ?thesis 
        using a0 seq_and_if_not_eq(12)[of Q] a0' ass last_length map_lift_catch_some_eq
        by (metis LanguageCon.com.distinct(17) LanguageCon.com.distinct(19))               
   } thus ?thesis using l ass by auto
qed


lemma not_func_redex_cptn_mod_nest_n': 
assumes a0:"Γ⊢c (P,s) → (Q, t)" and
        a1:"(n,Γ,(Q,t)#xs) ∈  cptn_mod_nest_call" and 
        a2:"(∀fn. redex P ≠ Call fn) ∨ 
            (redex P = Call fn ∧ Γ fn = None) ∨ 
            (redex P = Call fn ∧ (∀sa. s≠Normal sa))"
shows "(n,Γ,(P,s)#(Q,t)#xs) ∈  cptn_mod_nest_call"
using a0 a1 a2
proof (induct arbitrary: xs) 
  case (Basicc f s)   
  thus ?case by (simp add: Basicc cptn_mod_nest_call.CptnModNestSkip stepc.Basicc)
next
  case (Specc s t r)
  thus ?case by (simp add: Specc cptn_mod_nest_call.CptnModNestSkip stepc.Specc)
next
  case (SpecStuckc s r)
  thus ?case by (simp add: SpecStuckc cptn_mod_nest_call.CptnModNestSkip stepc.SpecStuckc)
next
  case (Guardc s g f c) 
    thus ?case by (simp add: cptn_mod_nest_call.CptnModNestGuard) 
next

  case (GuardFaultc s g f c)
  thus ?case by (simp add: GuardFaultc cptn_mod_nest_call.CptnModNestSkip stepc.GuardFaultc)
next
case (Seqc c1 s c1' s' c2)
  have step: "Γ⊢c (c1, s) → (c1', s')" by (simp add: Seqc.hyps(1))
  then have nsc1: "c1≠Skip" using stepc_elim_cases(1) by blast 
  have assum: "(n, Γ, (Seq c1' c2, s') # xs) ∈ cptn_mod_nest_call" using Seqc.prems by blast
  have divseq:"(∀s P Q zs. (Seq c1' c2, s') # xs=(Seq P Q, s)#zs ⟶
                (∃xs sv' sv''. ((n,Γ,(P, s)#xs) ∈ cptn_mod_nest_call  ∧ 
                           (zs=(map (lift Q) xs) ∨
                           ((fst(((P, s)#xs)!length xs)=Skip ∧ 
                             (∃ys. (n,Γ,(Q, snd(((P, s)#xs)!length xs))#ys) ∈ cptn_mod_nest_call ∧ 
                              zs=(map (lift (Q)) xs)@((Q, snd(((P, s)#xs)!length xs))#ys)))) ∨
                           ((fst(((P, s)#xs)!length xs)=Throw ∧ 
                               snd(last ((P, s)#xs)) = Normal sv' ∧  s'=Normal sv''∧
                             (∃ys.  (n,Γ,(Throw,Normal sv')#ys) ∈ cptn_mod_nest_call ∧
                              zs=(map (lift Q) xs)@((Throw,Normal sv')#ys))
                               ))))
                            
                 ))"  using div_seq_nest [OF assum] unfolding seq_cond_nest_def by auto
   {fix sa P Q zsa
       assume ass:"(Seq c1' c2, s') # xs = (Seq P Q, sa) # zsa"
       then have eqs:"c1' = P ∧ c2 = Q ∧ s' = sa ∧ xs = zsa" by auto
       then have "(∃xs sv' sv''. (n,Γ, (P, sa) # xs) ∈ cptn_mod_nest_call ∧
                        (zsa = map (lift Q) xs ∨              
                         fst (((P, sa) # xs) ! length xs) = Skip ∧
                             (∃ys. (n,Γ, (Q, snd (((P, sa) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
                             zsa = map (lift Q) xs @ (Q, snd (((P, sa) # xs) ! length xs)) # ys) ∨
                        ((fst(((P, sa)#xs)!length xs)=Throw ∧ 
                          snd(last ((P, sa)#xs)) = Normal sv' ∧  s'=Normal sv''∧
                          (∃ys.  (n,Γ,(Throw,Normal sv')#ys) ∈ cptn_mod_nest_call ∧
                              zsa=(map (lift Q) xs)@((Throw,Normal sv')#ys))))))" 
             using ass divseq by blast
    } note conc=this [of c1' c2 s' xs]    
     then obtain xs' sa' sa''
       where  split:"(n,Γ, (c1', s') # xs') ∈ cptn_mod_nest_call ∧
                     (xs = map (lift c2) xs' ∨                   
                     fst (((c1', s') # xs') ! length xs') = Skip ∧
                        (∃ys. (n,Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
                         xs = map (lift c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∨
                     ((fst(((c1', s')#xs')!length xs')=Throw ∧ 
                         snd(last ((c1', s')#xs')) = Normal sa' ∧ s'=Normal sa''∧
                         (∃ys.  (n,Γ,(Throw,Normal sa')#ys) ∈ cptn_mod_nest_call ∧
                              xs=(map (lift c2) xs')@((Throw,Normal sa')#ys))
                         )))"  by blast
  then have "(xs = map (lift c2) xs' ∨                   
                   fst (((c1', s') # xs') ! length xs') = Skip ∧
                      (∃ys. (n,Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
                       xs = map (lift c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∨
                   ((fst(((c1', s')#xs')!length xs')=Throw ∧ 
                       snd(last ((c1', s')#xs')) = Normal sa' ∧ s'=Normal sa''∧
                       (∃ys.  (n,Γ,(Throw,Normal sa')#ys) ∈ cptn_mod_nest_call ∧
                            xs=(map (lift c2) xs')@((Throw,Normal sa')#ys)))))" 
    by auto
  thus ?case 
  proof{           
       assume c1'nonf:"xs = map (lift c2) xs'"
       then have c1'cptn:"(n,Γ, (c1', s') # xs') ∈ cptn_mod_nest_call" using split by blast
       then have induct_step: "(n,Γ, (c1, s) # (c1', s')#xs') ∈ cptn_mod_nest_call"
         using Seqc.hyps(2) Seqc.prems(2) SmallStepCon.redex.simps(4) by auto  
       then have "(Seq c1' c2, s')#xs = map (lift c2) ((c1', s')#xs')"
         using c1'nonf
         by (simp add: lift_def)
       thus ?thesis 
         using c1'nonf c1'cptn induct_step by (auto simp add: CptnModNestSeq1)
    next
      assume "fst (((c1', s') # xs') ! length xs') = Skip ∧
              (∃ys. (n,Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
                  xs = map (lift c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∨
             ((fst(((c1', s')#xs')!length xs')=Throw ∧ 
                snd(last ((c1', s')#xs')) = Normal sa' ∧  s'=Normal sa''∧
                (∃ys.  (n,Γ,(Throw,Normal sa')#ys) ∈ cptn_mod_nest_call ∧
                              xs=(map (lift c2) xs')@((Throw,Normal sa')#ys))))"  
      thus ?thesis
      proof
       assume assth:"fst (((c1', s') # xs') ! length xs') = Skip ∧
            (∃ys. (n,Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
                xs = map (lift c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys)"
       then obtain ys 
           where split':"(n,Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
                xs = map (lift c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys"
           by auto    
       then have c1'cptn:"(n,Γ, (c1', s') # xs') ∈ cptn_mod_nest_call" using split by blast
       then have induct_step: "(n,Γ, (c1, s) # (c1', s')#xs') ∈ cptn_mod_nest_call"
       using Seqc.hyps(2) Seqc.prems(2) SmallStepCon.redex.simps(4) by auto                
       then have seqmap:"(Seq c1 c2, s)#(Seq c1' c2, s')#xs = map (lift c2) ((c1,s)#(c1', s')#xs') @ (c2, snd (((c1', s') # xs') ! length xs')) # ys"
      using split' by (simp add:  lift_def) 
      then have lastc1:"last ((c1, s) # (c1', s') # xs') = ((c1', s') # xs') ! length xs'"
        by (simp add: last_length) 
      then have lastc1skip:"fst (last ((c1, s) # (c1', s') # xs')) = Skip" 
           using assth by fastforce          
      thus ?thesis 
        using seqmap split' cptn_mod_nest_call.CptnModNestSeq2
              induct_step lastc1 lastc1skip
        by (metis (no_types) Cons_lift_append )               
    next
        assume assm:"((fst(((c1', s')#xs')!length xs')=Throw ∧ 
                snd(last ((c1', s')#xs')) = Normal sa' ∧  s'=Normal sa''∧
                (∃ys.(n,Γ,(Throw,Normal sa')#ys) ∈ cptn_mod_nest_call ∧
                 xs=(map (lift c2) xs')@((Throw,Normal sa')#ys))))"
        then have s'eqsa'': "s'=Normal sa''" by auto
        then have snormal: "∃ns. s=Normal ns" by (metis Seqc.hyps(1) step_Abrupt_prop step_Fault_prop step_Stuck_prop xstate.exhaust)
        then have c1'cptn:"(n,Γ, (c1', s') # xs') ∈ cptn_mod_nest_call" using split by blast        
        then have induct_step: "(n,Γ, (c1, s) # (c1', s')#xs') ∈ cptn_mod_nest_call"
        using Seqc.hyps(2) Seqc.prems(2) SmallStepCon.redex.simps(4) by auto 
        then obtain ys where seqmap:"(Seq c1' c2, s')#xs = (map (lift c2) ((c1', s')#xs'))@((Throw,Normal sa')#ys)"
        using assm
        proof -
          assume a1: "⋀ys. (LanguageCon.com.Seq c1' c2, s') # xs = map (lift c2) ((c1', s') # xs') @ (LanguageCon.com.Throw, Normal sa') # ys ⟹ thesis"
          have "(LanguageCon.com.Seq c1' c2, Normal sa'') # map (lift c2) xs' = map (lift c2) ((c1', s') # xs')"
            by (simp add: assm lift_def)
          thus ?thesis
            using a1 assm by moura
        qed  
        then have lastc1:"last ((c1, s) # (c1', s') # xs') = ((c1', s') # xs') ! length xs'"
                   by (simp add: last_length) 
        then have lastc1skip:"fst (last ((c1, s) # (c1', s') # xs')) = Throw" 
             using assm by fastforce
        then have "snd (last ((c1, s) # (c1', s') # xs')) = Normal sa'"
             using assm by force        
        thus ?thesis
            using assm c1'cptn induct_step lastc1skip snormal seqmap s'eqsa'' 
            by (auto simp add:cptn_mod_nest_call.CptnModNestSeq3)
   qed
  }qed    
next
  case (SeqSkipc c2 s xs)
  have c2incptn:"(n,Γ, (c2, s) # xs) ∈ cptn_mod_nest_call" by fact
  then have 1:"(n,Γ, [(Skip, s)]) ∈ cptn_mod_nest_call"
    by (simp add: cptn_mod_nest_call.CptnModNestOne)
  then have 2:"fst(last ([(Skip, s)])) = Skip" by fastforce
  then have 3:"(n,Γ,(c2, snd(last [(Skip, s)]))#xs) ∈ cptn_mod_nest_call" 
      using c2incptn by auto  
  then have "(c2,s)#xs=(map (lift c2) [])@(c2, snd(last [(Skip, s)]))#xs" 
       by (auto simp add:lift_def)
  thus ?case using 1 2 3 by (simp add: CptnModNestSeq2)   
next
  case (SeqThrowc c2 s xs)  
  have "(n,Γ, [(Throw, Normal s)]) ∈ cptn_mod_nest_call" 
    by (simp add: cptn_mod_nest_call.CptnModNestOne) 
  then obtain ys where 
    ys_nil:"ys=[]" and 
    last:"(n, Γ, (Throw, Normal s)#ys)∈ cptn_mod_nest_call"
   by auto
  moreover have "fst (last ((Throw, Normal s)#ys)) = Throw" using ys_nil last by auto
  moreover have "snd (last ((Throw, Normal s)#ys)) = Normal s" using ys_nil last by auto
  moreover from ys_nil have "(map (lift c2) ys) = []" by auto  
  ultimately show ?case using SeqThrowc.prems cptn_mod_nest_call.CptnModNestSeq3 by fastforce      

next
  case (CondTruec s b c1 c2)
  thus ?case by (simp add: cptn_mod_nest_call.CptnModNestCondT)
next
  case (CondFalsec s b c1 c2)
  thus ?case by (simp add: cptn_mod_nest_call.CptnModNestCondF)
next
 case (WhileTruec s1 b c)
 have sinb: "s1∈b" by fact
 have SeqcWhile: "(n,Γ, (Seq c (While b c), Normal s1) # xs) ∈ cptn_mod_nest_call" 
   by fact  
 have divseq:"(∀s P Q zs. (Seq c (While b c), Normal s1) # xs=(Seq P Q, s)#zs ⟶
                (∃xs s'. ((n,Γ,(P, s)#xs) ∈ cptn_mod_nest_call  ∧ 
                           (zs=(map (lift Q) xs) ∨
                           ((fst(((P, s)#xs)!length xs)=Skip ∧ 
                             (∃ys. (n,Γ,(Q, snd(((P, s)#xs)!length xs))#ys) ∈ cptn_mod_nest_call ∧ 
                              zs=(map (lift (Q)) xs)@((Q, snd(((P, s)#xs)!length xs))#ys)))) ∨
                           ((fst(((P, s)#xs)!length xs)=Throw ∧ 
                               snd(last ((P, s)#xs)) = Normal s' ∧
                              (∃ys.  (n,Γ,(Throw,Normal s')#ys) ∈ cptn_mod_nest_call ∧
                      zs=(map (lift Q) xs)@((Throw,Normal s')#ys))))))                            
                 ))" using div_seq_nest [OF SeqcWhile] by (auto simp add: seq_cond_nest_def)
{fix sa P Q zsa
       assume ass:"(Seq c (While b c), Normal s1) # xs = (Seq P Q, sa) # zsa"
       then have eqs:"c = P ∧ (While b c) = Q ∧ Normal s1 = sa ∧ xs = zsa" by auto
       then have "(∃xs s'. (n,Γ, (P, sa) # xs) ∈ cptn_mod_nest_call ∧
                        (zsa = map (lift Q) xs ∨              
                         fst (((P, sa) # xs) ! length xs) = Skip ∧
                             (∃ys. (n,Γ, (Q, snd (((P, sa) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
                             zsa = map (lift Q) xs @ (Q, snd (((P, sa) # xs) ! length xs)) # ys) ∨
                        ((fst(((P, sa)#xs)!length xs)=Throw ∧ 
                          snd(last ((P, sa)#xs)) = Normal s' ∧
                          (∃ys.  (n,Γ,(Throw,Normal s')#ys) ∈ cptn_mod_nest_call ∧
                      zsa=(map (lift Q) xs)@((Throw,Normal s')#ys))
                        ))))" 
             using ass divseq by auto
    } note conc=this [of c "While b c" "Normal s1" xs] 
   then obtain xs' s' 
        where  split:"(n,Γ, (c, Normal s1) # xs') ∈ cptn_mod_nest_call ∧
     (xs = map (lift (While b c)) xs' ∨
      fst (((c, Normal s1) # xs') ! length xs') = Skip ∧
      (∃ys. (n,Γ, (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys)
            ∈ cptn_mod_nest_call ∧
            xs =
            map (lift (While b c)) xs' @
            (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys) ∨
      fst (((c, Normal s1) # xs') ! length xs') = Throw ∧
      snd (last ((c, Normal s1) # xs')) = Normal s' ∧ 
      (∃ys.  (n,Γ, ((Throw, Normal s')#ys)) ∈ cptn_mod_nest_call ∧
      xs = map (lift (While b c)) xs' @ ((Throw, Normal s')#ys)))"  by auto
 then have "(xs = map (lift (While b c)) xs' ∨
            fst (((c, Normal s1) # xs') ! length xs') = Skip ∧
            (∃ys. (n,Γ, (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys)
                  ∈ cptn_mod_nest_call ∧
                  xs =
                  map (lift (While b c)) xs' @
                  (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys) ∨
            fst (((c, Normal s1) # xs') ! length xs') = Throw ∧
            snd (last ((c, Normal s1) # xs')) = Normal s' ∧
            (∃ys.  (n,Γ, ((Throw, Normal s')#ys)) ∈ cptn_mod_nest_call ∧
          xs = map (lift (While b c)) xs' @ ((Throw, Normal s')#ys)))" ..
 thus ?case
 proof{ 
   assume 1:"xs = map (lift (While b c)) xs'"  
   have 3:"(n, Γ, (c, Normal s1) # xs') ∈ cptn_mod_nest_call" using split by auto   
   then show ?thesis 
     using "1" cptn_mod_nest_call.CptnModNestWhile1 sinb by fastforce 
 next
   assume "fst (((c, Normal s1) # xs') ! length xs') = Skip ∧
          (∃ys. (n,Γ, (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys)
                ∈ cptn_mod_nest_call ∧
                xs =
                map (lift (While b c)) xs' @
                (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys) ∨
          fst (((c, Normal s1) # xs') ! length xs') = Throw ∧
          snd (last ((c, Normal s1) # xs')) = Normal s' ∧
          (∃ys.  (n,Γ, ((Throw, Normal s')#ys)) ∈ cptn_mod_nest_call ∧
          xs = map (lift (While b c)) xs' @ ((Throw, Normal s')#ys))"
   thus ?case
   proof
     assume asm:"fst (((c, Normal s1) # xs') ! length xs') = Skip ∧
             (∃ys. (n,Γ, (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys)
             ∈ cptn_mod_nest_call ∧
             xs =
             map (lift (While b c)) xs' @
             (While b c, snd (((c, Normal s1) # xs') ! length xs')) # ys)"
     then obtain ys 
       where asm':"(n,Γ, (While b c, snd (last ((c, Normal s1) # xs'))) # ys)
                   ∈ cptn_mod_nest_call 
                   ∧ xs = map (lift (While b c)) xs' @
                       (While b c, snd (last ((c, Normal s1) # xs'))) # ys" 
              by (auto simp add: last_length)
     moreover have 3:"(n,Γ, (c, Normal s1) # xs') ∈ cptn_mod_nest_call" using split by auto
     moreover from asm have "fst (last ((c, Normal s1) # xs'))  = Skip"
          by (simp add: last_length) 
     ultimately show ?case using sinb by (auto simp add:CptnModNestWhile2)
   next
    assume asm:" fst (((c, Normal s1) # xs') ! length xs') = Throw ∧
          snd (last ((c, Normal s1) # xs')) = Normal s' ∧
          (∃ys.  (n,Γ, ((Throw, Normal s')#ys)) ∈ cptn_mod_nest_call ∧
          xs = map (lift (While b c)) xs' @ ((Throw, Normal s')#ys))"   
     moreover have 3:"(n,Γ, (c, Normal s1) # xs') ∈ cptn_mod_nest_call" 
       using split by auto
     moreover from asm have "fst (last ((c, Normal s1) # xs'))  = Throw"
          by (simp add: last_length) 
     ultimately show ?case using sinb by (auto simp add:CptnModNestWhile3)
   qed
 }qed  
next
 case (WhileFalsec s b c)
 thus ?case by (simp add: cptn_mod_nest_call.CptnModNestSkip stepc.WhileFalsec)
next
  case (Awaitc s b c t)
  thus ?case by (simp add: cptn_mod_nest_call.CptnModNestSkip stepc.Awaitc)
next
  case (AwaitAbruptc s b c t t')
  thus ?case by (simp add: cptn_mod_nest_call.CptnModNestThrow stepc.AwaitAbruptc) 
next
  case (Callc p bdy s)
  thus ?case using SmallStepCon.redex.simps(7) by auto  
next
  case (CallUndefinedc p s)
  then have "p = fn" by auto
  thus ?case using CallUndefinedc
  proof -
    have "(LanguageCon.com.Call fn ∩gs (LanguageCon.com.Skip::('b, 'a, 'c,'d) LanguageCon.com)) ≠ Some LanguageCon.com.Skip"
      by simp
    then show ?thesis
      by (metis (no_types) CallUndefinedc.hyps LanguageCon.com.inject(6) LanguageCon.inter_guards.simps(79) SmallStepCon.redex.simps(7) ‹(n, Γ, (LanguageCon.com.Skip, Stuck) # xs) ∈ cptn_mod_nest_call› cptn_mod_nest_call.CptnModNestSkip stepc.CallUndefinedc)
  qed 
next
  case (DynComc c s)
  thus ?case by (simp add: cptn_mod_nest_call.CptnModNestDynCom) 
next
  case (Catchc c1 s c1' s' c2)
   have step: "Γ⊢c (c1, s) → (c1', s')" by (simp add: Catchc.hyps(1))
  then have nsc1: "c1≠Skip" using stepc_elim_cases(1) by blast 
  have assum: "(n,Γ, (Catch c1' c2, s') # xs) ∈ cptn_mod_nest_call" 
  using Catchc.prems by blast
  have divcatch:"(∀s P Q zs. (Catch c1' c2, s') # xs=(Catch P Q, s)#zs ⟶
  (∃xs s' s''. ((n,Γ,(P, s)#xs) ∈ cptn_mod_nest_call  ∧ 
             (zs=(map (lift_catch Q) xs) ∨
             ((fst(((P, s)#xs)!length xs)=Throw ∧
               snd(last ((P, s)#xs)) = Normal s' ∧  s=Normal s''∧
               (∃ys. (n,Γ,(Q, snd(((P, s)#xs)!length xs))#ys) ∈ cptn_mod_nest_call ∧ 
                zs=(map (lift_catch Q) xs)@((Q, snd(((P, s)#xs)!length xs))#ys)))) ∨
                ((fst(((P, s)#xs)!length xs)=Skip ∧  
                  (∃ys. (n,Γ,(Skip,snd(last ((P, s)#xs)))#ys) ∈ cptn_mod_nest_call ∧                   
                 zs=(map (lift_catch Q) xs)@((Skip,snd(last ((P, s)#xs)))#ys))                
                 ))))
   ))" using div_catch_nest [OF assum] by (auto simp add: catch_cond_nest_def)
   {fix sa P Q zsa
       assume ass:"(Catch c1' c2, s') # xs = (Catch P Q, sa) # zsa"
       then have eqs:"c1' = P ∧ c2 = Q ∧ s' = sa ∧ xs = zsa" by auto
       then have "(∃xs sv' sv''. ((n, Γ,(P, sa)#xs) ∈ cptn_mod_nest_call  ∧ 
             (zsa=(map (lift_catch Q) xs) ∨
             ((fst(((P, sa)#xs)!length xs)=Throw ∧
               snd(last ((P, sa)#xs)) = Normal sv' ∧  s'=Normal sv''∧
               (∃ys. (n,Γ,(Q, snd(((P, sa)#xs)!length xs))#ys) ∈ cptn_mod_nest_call ∧ 
                zsa=(map (lift_catch Q) xs)@((Q, snd(((P, sa)#xs)!length xs))#ys)))) ∨
                ((fst(((P, sa)#xs)!length xs)=Skip ∧                  
                 (∃ys. (n,Γ,(Skip,snd(last ((P, sa)#xs)))#ys) ∈ cptn_mod_nest_call ∧                   
                 zsa=(map (lift_catch Q) xs)@((Skip,snd(last ((P, sa)#xs)))#ys))))))
   )"   using ass divcatch by blast
    } note conc=this [of c1' c2 s' xs]    
     then obtain xs' sa' sa''
       where split:
         "(n,Γ, (c1', s') # xs') ∈ cptn_mod_nest_call ∧ 
          (xs = map (lift_catch c2) xs' ∨
          fst (((c1', s') # xs') ! length xs') = Throw ∧
          snd (last ((c1', s') # xs')) = Normal sa' ∧ s' = Normal sa'' ∧
          (∃ys. (n,Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
                xs = map (lift_catch c2) xs' @
                (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∨
          fst (((c1', s') # xs') ! length xs') = Skip ∧
          (∃ys. (n,Γ,(Skip,snd(last ((c1', s')#xs')))#ys) ∈ cptn_mod_nest_call ∧                   
                 xs=(map (lift_catch c2) xs')@((Skip,snd(last ((c1', s')#xs')))#ys)))"          
        by blast
  then have "(xs = map (lift_catch c2) xs' ∨
          fst (((c1', s') # xs') ! length xs') = Throw ∧
          snd (last ((c1', s') # xs')) = Normal sa' ∧ s' = Normal sa'' ∧
          (∃ys. (n,Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
                xs = map (lift_catch c2) xs' @
                (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∨
          fst (((c1', s') # xs') ! length xs') = Skip ∧
          (∃ys. (n,Γ,(Skip,snd(last ((c1', s')#xs')))#ys) ∈ cptn_mod_nest_call ∧                   
                 xs=(map (lift_catch c2) xs')@((Skip,snd(last ((c1', s')#xs')))#ys)))"          
        by auto
  thus ?case 
  proof{           
       assume c1'nonf:"xs = map (lift_catch c2) xs'"
       then have c1'cptn:"(n,Γ, (c1', s') # xs') ∈ cptn_mod_nest_call" using split by blast
       then have induct_step: "(n, Γ, (c1, s) # (c1', s')#xs') ∈ cptn_mod_nest_call"
         using Catchc.hyps(2) Catchc.prems(2) SmallStepCon.redex.simps(11) by auto  
       then have "(Catch c1' c2, s')#xs = map (lift_catch c2) ((c1', s')#xs')"
            using c1'nonf
            by (simp add: CptnModCatch1 lift_catch_def)
       thus ?thesis 
            using c1'nonf c1'cptn induct_step 
       by (auto simp add: CptnModNestCatch1)
    next
      assume "fst (((c1', s') # xs') ! length xs') = Throw ∧
                snd (last ((c1', s') # xs')) = Normal sa' ∧ s' = Normal sa'' ∧
                (∃ys. (n,Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
                xs =map (lift_catch c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∨
               fst (((c1', s') # xs') ! length xs') = Skip ∧
                (∃ys. (n,Γ,(Skip,snd(last ((c1', s')#xs')))#ys) ∈ cptn_mod_nest_call ∧                   
                 xs=(map (lift_catch c2) xs')@((Skip,snd(last ((c1', s')#xs')))#ys))"  
      thus ?thesis
      proof
         assume assth:
               "fst (((c1', s') # xs') ! length xs') = Throw ∧
                snd (last ((c1', s') # xs')) = Normal sa' ∧ s' = Normal sa'' ∧
                (∃ys. (n,Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
                xs =map (lift_catch c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys)"
             then have s'eqsa'': "s'=Normal sa''" by auto
             then have snormal: "∃ns. s=Normal ns" by (metis Catchc.hyps(1) step_Abrupt_prop step_Fault_prop step_Stuck_prop xstate.exhaust)
             then obtain ys 
             where split':"(n,Γ, (c2, snd (((c1', s') # xs') ! length xs')) # ys) ∈ cptn_mod_nest_call ∧
                xs =map (lift_catch c2) xs' @ (c2, snd (((c1', s') # xs') ! length xs')) # ys"
                using assth by auto    
         then have c1'cptn:"(n,Γ, (c1', s') # xs') ∈ cptn_mod_nest_call" 
              using split by blast
         then have induct_step: "(n,Γ, (c1, s) # (c1', s')#xs') ∈ cptn_mod_nest_call"
              using Catchc.hyps(2) Catchc.prems(2) SmallStepCon.redex.simps(11) by auto                
         then have seqmap:"(Catch c1 c2, s)#(Catch c1' c2, s')#xs = map (lift_catch c2) ((c1,s)#(c1', s')#xs') @ (c2, snd (((c1', s') # xs') ! length xs')) # ys"
              using split' by (simp add: CptnModCatch3 lift_catch_def) 
        then have lastc1:"last ((c1, s) # (c1', s') # xs') = ((c1', s') # xs') ! length xs'"
             by (simp add: last_length) 
        then have lastc1skip:"fst (last ((c1, s) # (c1', s') # xs')) = Throw" 
             using assth by fastforce
        then have "snd (last ((c1, s) # (c1', s') # xs')) = Normal sa'"
             using assth by force
        thus ?thesis using snormal seqmap s'eqsa'' split' 
              last_length  cptn_mod_nest_call.CptnModNestCatch3 
              induct_step lastc1 lastc1skip
              using Cons_lift_catch_append by fastforce           
    next
        assume assm:" fst (((c1', s') # xs') ! length xs') = Skip ∧ 
                       (∃ys. (n,Γ,(Skip,snd(last ((c1', s')#xs')))#ys) ∈ cptn_mod_nest_call ∧                   
                      xs=(map (lift_catch c2) xs')@((Skip,snd(last ((c1', s')#xs')))#ys))"
        then have c1'cptn:"(n,Γ, (c1', s') # xs') ∈ cptn_mod_nest_call" using split by blast
        then have induct_step: "(n,Γ, (c1, s) # (c1', s')#xs') ∈ cptn_mod_nest_call"
        using Catchc.hyps(2) Catchc.prems(2) SmallStepCon.redex.simps(11) by auto 
        then have "map (lift_catch c2) ((c1', s') # xs') = (Catch c1' c2, s') # map (lift_catch c2) xs'" 
          by (auto simp add: lift_catch_def)
        then obtain ys 
             where seqmap:"(Catch c1' c2, s')#xs = (map (lift_catch c2) ((c1', s')#xs'))@((Skip,snd(last ((c1', s')#xs')))#ys)"
        using assm by fastforce
        then have lastc1:"last ((c1, s) # (c1', s') # xs') = ((c1', s') # xs') ! length xs'"
                   by (simp add: last_length) 
        then have lastc1skip:"fst (last ((c1, s) # (c1', s') # xs')) = Skip" 
             using assm by fastforce
        then have "snd (last ((c1, s) # (c1', s') # xs')) = snd (last ((c1', s') # xs'))"
             using assm by force
        thus ?thesis 
            using assm c1'cptn induct_step lastc1skip seqmap 
            by (auto simp add:cptn_mod_nest_call.CptnModNestCatch2)
    qed
  }qed              
next
  case (CatchThrowc c2 s)
  have c2incptn:"(n,Γ, (c2, Normal s) # xs) ∈ cptn_mod_nest_call" by fact
  then have 1:"(n,Γ, [(Throw, Normal s)]) ∈ cptn_mod_nest_call" 
    by (simp add: cptn_mod_nest_call.CptnModNestOne)
  then have 2:"fst(last ([(Throw, Normal s)])) = Throw" by fastforce
  then have 3:"(n,Γ,(c2, snd(last [(Throw, Normal s)]))#xs) ∈ cptn_mod_nest_call" 
      using c2incptn by auto  
  then have "(c2,Normal s)#xs=(map (lift c2) [])@(c2, snd(last [(Throw, Normal s)]))#xs" 
       by (auto simp add:lift_def)
  thus ?case using 1 2 3 by (simp add: CptnModNestCatch3)
next
  case (CatchSkipc c2 s)
  have "(n,Γ, [(Skip, s)]) ∈ cptn_mod_nest_call" 
    by (simp add: cptn_mod_nest_call.CptnModNestOne)
  then obtain ys where 
    ys_nil:"ys=[]" and 
    last:"(n,Γ, (Skip,  s)#ys)∈ cptn_mod_nest_call"
    by auto
  moreover have "fst (last ((Skip,  s)#ys)) = Skip" using ys_nil last by auto
  moreover have "snd (last ((Skip,  s)#ys)) = s" using ys_nil last by auto
  moreover from ys_nil have "(map (lift_catch c2) ys) = []" by auto  
  ultimately show ?case using CatchSkipc.prems 
     by simp (simp add: cptn_mod_nest_call.CptnModNestCatch2 ys_nil)
next
  case (FaultPropc c f)
  thus ?case by (simp add: cptn_mod_nest_call.CptnModNestSkip stepc.FaultPropc) 
next
  case (AbruptPropc c f)
  thus ?case by (simp add: cptn_mod_nest_call.CptnModNestSkip stepc.AbruptPropc)
next
  case (StuckPropc c)
  thus ?case by (simp add: cptn_mod_nest_call.CptnModNestSkip stepc.StuckPropc)
qed




lemma not_func_redex_cptn_mod_nest_n_env: 
assumes a0:"Γ⊢c (P,s) →e (P, t)" and
        a1:"(n,Γ,(P,t)#xs) ∈  cptn_mod_nest_call"         
shows "(n,Γ,(P,s)#(P,t)#xs) ∈  cptn_mod_nest_call"
  by (simp add: a0 a1 cptn_mod_nest_call.CptnModNestEnv)


lemma cptn_mod_nest_cptn_mod:"(n,Γ,cfs) ∈  cptn_mod_nest_call  ⟹ (Γ,cfs)∈ cptn_mod"
by (induct rule:cptn_mod_nest_call.induct, (fastforce simp:cptn_mod.intros )+)


lemma cptn_mod_cptn_mod_nest: "(Γ,cfs)∈ cptn_mod ⟹ ∃n. (n,Γ,cfs) ∈  cptn_mod_nest_call"
proof (induct rule:cptn_mod.induct)
   case (CptnModSkip Γ P s t xs) 
   then obtain n where cptn_nest:"(n, Γ, (Skip, t) # xs) ∈ cptn_mod_nest_call" by auto   
    {assume asm:"∀f. ((∃sn. s = Normal sn) ∧ (Γ f) = Some Skip ⟶ P  ≠ Call f  )"     
     then have ?case using CptnModNestSkip[OF CptnModSkip(1) CptnModSkip(2) asm cptn_nest] by auto
    }note t1=this
    {assume asm:"¬ (∀f. ((∃sn. s = Normal sn) ∧ (Γ f) = Some Skip ⟶ P  ≠ Call f))"
     then obtain f where asm:"((∃sn. s = Normal sn) ∧ (Γ f) = Some Skip ∧ P  = Call f)" by auto 
      then obtain sn where normal_s:"s=Normal sn" by auto
     then have t_eq_s:"t=s" using asm cptn_nest normal_s
       by (metis CptnModSkip.hyps(1) LanguageCon.com.simps(22) 
           LanguageCon.inter_guards.simps(79) LanguageCon.inter_guards_Call 
           Pair_inject stepc_Normal_elim_cases(9))
     then have "(Suc n, Γ,((Call f), Normal sn)#(Skip, Normal sn)#xs) ∈ cptn_mod_nest_call"
       using asm cptn_nest normal_s CptnModNestCall by fastforce        
     then have ?case using asm normal_s t_eq_s by fastforce
    }note t2 = this
    then show ?case using t1 t2 by fastforce  
next
   case (CptnModThrow Γ P s t xs)  
   then obtain n where cptn_nest:"(n, Γ, (Throw, t) # xs) ∈ cptn_mod_nest_call" by auto   
    {assume asm:"∀f. ((∃sn. s = Normal sn) ∧ (Γ f) = Some Throw ⟶ P  ≠ Call f  )"     
     then have ?case using CptnModNestThrow[OF CptnModThrow(1) CptnModThrow(2) asm cptn_nest] by auto
    }note t1=this
    {assume asm:"¬ (∀f. ((∃sn. s = Normal sn) ∧ (Γ f) = Some Throw ⟶ P  ≠ Call f))"
     then obtain f where asm:"((∃sn. s = Normal sn) ∧ (Γ f) = Some Throw ∧ P  = Call f)" by auto 
      then obtain sn where normal_s:"s=Normal sn" by auto
     then have t_eq_s:"t=s" using asm cptn_nest normal_s
       by (metis CptnModThrow.hyps(1) LanguageCon.com.simps(22) 
           LanguageCon.inter_guards.simps(79) LanguageCon.inter_guards_Call 
           Pair_inject stepc_Normal_elim_cases(9))
     then have "(Suc n, Γ,((Call f), Normal sn)#(Throw, Normal sn)#xs) ∈ cptn_mod_nest_call"
       using asm cptn_nest normal_s CptnModNestCall by fastforce        
     then have ?case using asm normal_s t_eq_s by fastforce
    }note t2 = this
    then show ?case using t1 t2 by fastforce
next
   case (CptnModSeq2 Γ P0 s xs P1 ys zs) 
   obtain n where n:"(n, Γ, (P0, s) # xs) ∈ cptn_mod_nest_call" using CptnModSeq2(2) by auto
   also obtain m where m:"(m, Γ, (P1, snd (last ((P0, s) # xs))) # ys) ∈ cptn_mod_nest_call"
     using CptnModSeq2(5) by auto
   ultimately show ?case
   proof (cases "n≥m") 
     case True thus ?thesis  
       using cptn_mod_nest_mono[of m Γ _ n] m n CptnModSeq2 cptn_mod_nest_call.CptnModNestSeq2 by blast
   next  
     case False 
     thus ?thesis
       using cptn_mod_nest_mono[of n Γ _ m] m n CptnModSeq2 
             cptn_mod_nest_call.CptnModNestSeq2 le_cases3 by blast      
   qed
next
   case (CptnModSeq3 Γ P0 s xs s' ys zs P1) 
   obtain n where n:"(n, Γ, (P0, Normal s) # xs) ∈ cptn_mod_nest_call" using CptnModSeq3(2) by auto
   also obtain m where m:"(m, Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn_mod_nest_call"
     using CptnModSeq3(6) by auto
   ultimately show ?case
   proof (cases "n≥m") 
     case True thus ?thesis  
       using cptn_mod_nest_mono[of m Γ _ n] m n CptnModSeq3 cptn_mod_nest_call.CptnModNestSeq3
       by fastforce
   next  
     case False 
     thus ?thesis
       using cptn_mod_nest_mono[of n Γ _ m] m n CptnModSeq3
             cptn_mod_nest_call.CptnModNestSeq3 le_cases3
      proof -
        have f1: "¬ n ≤ m ∨ (m, Γ, (P0, Normal s) # xs) ∈ cptn_mod_nest_call"
          by (metis cptn_mod_nest_mono[of n Γ _ m] n)
        have "n ≤ m"
          using False by linarith
        then have "(m, Γ, (P0, Normal s) # xs) ∈ cptn_mod_nest_call"
          using f1 by metis
        then show ?thesis
          by (metis (no_types) CptnModSeq3(3) CptnModSeq3(4) CptnModSeq3(7) 
                   cptn_mod_nest_call.CptnModNestSeq3 m)
      qed          
   qed
next
   case (CptnModWhile2 Γ P s xs b zs ys) 
   obtain n where n:"(n, Γ, (P, Normal s) # xs) ∈ cptn_mod_nest_call" using CptnModWhile2(2) by auto
   also obtain m where 
     m:" (m, Γ, (LanguageCon.com.While b P, snd (last ((P, Normal s) # xs))) # ys) ∈ 
          cptn_mod_nest_call"
     using CptnModWhile2(7) by auto
   ultimately show ?case
   proof (cases "n≥m") 
     case True thus ?thesis  
       using cptn_mod_nest_mono[of m Γ _ n] m n 
             CptnModWhile2 cptn_mod_nest_call.CptnModNestWhile2 by metis
   next  
     case False 
     thus ?thesis       
    proof -
      have f1: "¬ n ≤ m ∨ (m, Γ, (P, Normal s) # xs) ∈ cptn_mod_nest_call"
        using cptn_mod_nest_mono[of n Γ _ m] n by presburger
      have "n ≤ m"
        using False by linarith
      then have "(m, Γ, (P, Normal s) # xs) ∈ cptn_mod_nest_call"
        using f1 by metis
      then show ?thesis
        by (metis (no_types) CptnModWhile2(3) CptnModWhile2(4) CptnModWhile2(5) 
                  cptn_mod_nest_call.CptnModNestWhile2 m)
    qed 
   qed
next
   case (CptnModWhile3 Γ P s xs b s' ys zs)
   obtain n where n:"(n, Γ, (P, Normal s) # xs) ∈ cptn_mod_nest_call" 
     using CptnModWhile3(2) by auto
   also obtain m where 
     m:" (m, Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn_mod_nest_call"
     using CptnModWhile3(7) by auto
   ultimately show ?case
   proof (cases "n≥m") 
     case True thus ?thesis  
     proof -
      have "(n, Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn_mod_nest_call"
        using True cptn_mod_nest_mono[of m Γ _ n] m by presburger
      then show ?thesis
        by (metis (no_types) CptnModWhile3.hyps(3) CptnModWhile3.hyps(4) 
            CptnModWhile3.hyps(5) CptnModWhile3.hyps(8) cptn_mod_nest_call.CptnModNestWhile3 n)
     qed 
   next  
     case False 
     thus ?thesis  using m n cptn_mod_nest_call.CptnModNestWhile3 cptn_mod_nest_mono[of n Γ _ m]
       by (metis CptnModWhile3.hyps(3) CptnModWhile3.hyps(4) 
           CptnModWhile3.hyps(5) CptnModWhile3.hyps(8) le_cases)
   qed
next
  case (CptnModCatch2 Γ P0 s xs ys zs P1) 
  obtain n where n:"(n, Γ, (P0, s) # xs) ∈ cptn_mod_nest_call" using CptnModCatch2(2) by auto
   also obtain m where m:"(m, Γ, (LanguageCon.com.Skip, snd (last ((P0, s) # xs))) # ys) ∈ cptn_mod_nest_call"
     using CptnModCatch2(5) by auto
   ultimately show ?case
   proof (cases "n≥m") 
     case True thus ?thesis  
       using cptn_mod_nest_mono[of m Γ _ n] m n 
             CptnModCatch2 cptn_mod_nest_call.CptnModNestCatch2 by blast
   next  
     case False 
     thus ?thesis
       using cptn_mod_nest_mono[of n Γ _ m] m n CptnModCatch2 
             cptn_mod_nest_call.CptnModNestCatch2 le_cases3 by blast      
   qed
next
  case (CptnModCatch3 Γ P0 s xs s' ys zs P1) 
   obtain n where n:"(n, Γ, (P0, Normal s) # xs) ∈ cptn_mod_nest_call" 
     using CptnModCatch3(2) by auto
   also obtain m where m:"(m, Γ, (ys, snd (last ((P0, Normal s) # xs))) # zs) ∈ cptn_mod_nest_call"
     using CptnModCatch3(6) by auto
   ultimately show ?case
   proof (cases "n≥m") 
     case True thus ?thesis  
       using cptn_mod_nest_mono[of m Γ _ n] m n CptnModCatch3 cptn_mod_nest_call.CptnModNestCatch3
       by fastforce
   next  
     case False 
     thus ?thesis
       using cptn_mod_nest_mono[of n Γ _ m] m n CptnModCatch3 
             cptn_mod_nest_call.CptnModNestCatch3 le_cases3
      proof -
        have f1: "¬ n ≤ m ∨ (m, Γ, (P0, Normal s) # xs) ∈ cptn_mod_nest_call"
          using ‹⋀cfs. ⟦(n, Γ, cfs) ∈ cptn_mod_nest_call; n ≤ m⟧ ⟹ (m, Γ, cfs) ∈ cptn_mod_nest_call› n by presburger
        have "n ≤ m"
          using False by auto
        then have "(m, Γ, (P0, Normal s) # xs) ∈ cptn_mod_nest_call"
          using f1 by meson
        then show ?thesis
          by (metis (no_types) ‹P1 = map (lift_catch ys) xs @ (ys, snd (last ((P0, Normal s) # xs))) # zs› ‹fst (last ((P0, Normal s) # xs)) = LanguageCon.com.Throw› ‹snd (last ((P0, Normal s) # xs)) = Normal s'› cptn_mod_nest_call.CptnModNestCatch3 m)
      qed      
   qed
qed(fastforce intro: cptn_mod_nest_call.intros)+

lemma cptn_mod_eq_cptn_mod_nest:
  "(Γ,cfs)∈ cptn_mod ⟷ (∃n. (n,Γ,cfs) ∈  cptn_mod_nest_call)"
  using cptn_mod_cptn_mod_nest cptn_mod_nest_cptn_mod by auto

lemma cptn_mod_eq_cptn_mod_nest':
  "∃n. ((Γ,cfs)∈ cptn_mod ⟷ (n,Γ,cfs) ∈  cptn_mod_nest_call)"
  using cptn_mod_eq_cptn_mod_nest by auto

lemma cptn_mod_eq_cptn_mod_nest1:
  "(Γ,cfs)∈ cptn_mod = (∃n. (n,Γ,cfs) ∈  cptn_mod_nest_call)"
  using cptn_mod_cptn_mod_nest cptn_mod_nest_cptn_mod by auto


lemma cptn_eq_cptn_mod_nest:
  "(Γ,cfs)∈ cptn = (∃n. (n,Γ,cfs) ∈  cptn_mod_nest_call)"
  using cptn_eq_cptn_mod_set cptn_mod_cptn_mod_nest cptn_mod_nest_cptn_mod by blast


subsection ‹computation on nested calls limit›
subsection ‹Elimination theorems›
lemma mod_env_not_component:
shows    "¬ Γ⊢c (P, s) → (P, t)"
proof 
  assume a3:"Γ⊢c (P, s) → (P, t)"           
  thus  False using step_change_p_or_eq_s a3 by fastforce
qed

lemma elim_cptn_mod_nest_step_c:
 assumes a0:"(n,Γ,cfg) ∈  cptn_mod_nest_call" and
         a1:"cfg = (P,s)#(Q,t)#cfg1"         
 shows "Γ⊢c (P,s) → (Q,t) ∨ Γ⊢c (P,s) →e (Q,t)"
proof-
  have "(Γ,cfg) ∈  cptn" using a0 cptn_mod_nest_cptn_mod
    using cptn_eq_cptn_mod_set by auto 
  then have "Γ⊢c (P,s) →ce (Q,t)" using a1
    by (metis c_step cptn_elim_cases(2) e_step)
  thus ?thesis
    using step_ce_not_step_e_step_c by blast  
qed

lemma elim_cptn_mod_nest_call_env:
 assumes a0:"(n,Γ,cfg) ∈  cptn_mod_nest_call" and
         a1:"cfg = (P,s)#(P,t)#cfg1"  and
         a2:"∀f. Γ f = Some (LanguageCon.com.Call f) ∧ 
                 (∃sn. s = Normal sn) ∧ s = t ⟶ SmallStepCon.redex P ≠ LanguageCon.com.Call f"
 shows "(n,Γ,(P,t)#cfg1) ∈  cptn_mod_nest_call"
 using a0 a1 a2
proof (induct arbitrary: P cfg1 s t rule:cptn_mod_nest_call.induct ) 
case (CptnModNestSeq1 n Γ P0 sa xs zs P1)
   then obtain xs' where "xs =  (P0, t)#xs'" unfolding lift_def by fastforce
   then have step:"(n, Γ, (P0, t) # xs') ∈ cptn_mod_nest_call" using CptnModNestSeq1 by fastforce        
   have "(P, t) = lift P1 (P0, t) ∧ cfg1 = map (lift P1) xs'"
      using CptnModNestSeq1.hyps(3) CptnModNestSeq1.prems(1) ‹xs = (P0, t) # xs'› by auto
    then have "(n, Γ, (LanguageCon.com.Seq P0 P1, t) # cfg1) ∈ cptn_mod_nest_call"
      by (meson cptn_mod_nest_call.CptnModNestSeq1 local.step)
    then show ?case
      using CptnModNestSeq1.prems(1) by fastforce  
next
  case (CptnModNestSeq2 n Γ P0 sa xs P1 ys zs) 
  thus ?case 
  proof (induct xs)
    case Nil thus ?case using Nil.prems(6) Nil.prems(7) by force
  next
    case (Cons x xs') 
    then have x:"x=(P0,t)" 
    proof-
      have "zs=(Seq P0 P1,t)#cfg1" using Cons by fastforce
      thus ?thesis using Cons(7) unfolding lift_def
      proof -
        assume "zs = map (λa. case a of (P, s) ⇒ (LanguageCon.com.Seq P P1, s)) (x # xs') @ 
                     (P1, snd (last ((P0, sa) # x # xs'))) # ys"
        then have "LanguageCon.com.Seq (fst x) P1 = LanguageCon.com.Seq P0 P1 ∧ snd x = t"
          by (simp add: ‹zs = (LanguageCon.com.Seq P0 P1, t) # cfg1› case_prod_beta)
        then show ?thesis
          by fastforce
      qed 
    qed
    then have step:"(n, Γ, (P0, t) # xs') ∈ cptn_mod_nest_call" using Cons by fastforce         
    have "fst (last ((P0, t) # xs')) = LanguageCon.com.Skip"
      using Cons.prems(3) ‹x = (P0, t)› by force
    then show ?case
      using Cons.prems(4) Cons.prems(6) CptnModNestSeq2.prems(1) x 
            cptn_mod_nest_call.CptnModNestSeq2 local.step by fastforce
  qed          
next 
  case (CptnModNestSeq3 n Γ P0 sa xs s' ys zs P1) 
  thus ?case 
  proof (induct xs)
    case Nil thus ?case using Nil.prems(6) Nil.prems(7) by force
  next
    case (Cons x xs') 
    then have x:"x=(P0,t)" 
    proof-
      have zs:"zs=(Seq P0 P1,t)#cfg1" using Cons by fastforce
      have "(LanguageCon.com.Seq (fst x) P1, snd x) = lift P1 x"
         by (simp add: lift_def prod.case_eq_if)
      then have "LanguageCon.com.Seq (fst x) P1 = LanguageCon.com.Seq P0 P1 ∧ snd x = t"
         using Cons.prems(7) zs by force
      then show ?thesis
          by fastforce                
    qed
    then have step:"(n, Γ, (P0, t) # xs') ∈ cptn_mod_nest_call" using Cons by fastforce         
    then obtain t' where t:"t=Normal t'" 
      using Normal_Normal Cons(2) Cons(5) cptn_mod_nest_cptn_mod cptn_eq_cptn_mod_set x
      by (metis snd_eqD)
    then show ?case using x Cons(5) Cons(6) cptn_mod_nest_call.CptnModNestSeq3 step
    proof -
      have "last ((P0, Normal t') # xs') = last ((P0, Normal sa) # x # xs')"
        using t x by force
      then have "fst (last ((P0, Normal t') # xs')) = LanguageCon.com.Throw"
        using Cons.prems(3) by presburger
      then show ?thesis
        using Cons.prems(4) Cons.prems(5) Cons.prems(7) 
              CptnModNestSeq3.prems(1) cptn_mod_nest_call.CptnModNestSeq3 
              local.step t x by fastforce
    qed       
  qed       
next
  case (CptnModNestCatch1 n Γ P0 s xs zs P1) 
   then obtain xs' where "xs =  (P0, t)#xs'" unfolding lift_catch_def by fastforce
   then have step:"(n, Γ, (P0, t) # xs') ∈ cptn_mod_nest_call" using CptnModNestCatch1 by fastforce        
   have "(P, t) = lift_catch P1 (P0, t) ∧ cfg1 = map (lift_catch P1) xs'"
      using CptnModNestCatch1.hyps(3) CptnModNestCatch1.prems(1) ‹xs = (P0, t) # xs'› by auto
    then have "(n, Γ, (Catch P0 P1, t) # cfg1) ∈ cptn_mod_nest_call"
      by (meson cptn_mod_nest_call.CptnModNestCatch1 local.step)
    then show ?case
      using CptnModNestCatch1.prems(1) by fastforce  
next
  case (CptnModNestCatch2 n Γ P0 sa xs ys zs P1) 
  thus ?case 
  proof (induct xs)
    case Nil thus ?case using Nil.prems(6) Nil.prems(7) by force
  next
    case (Cons x xs') 
    then have x:"x=(P0,t)" 
    proof-
      have zs:"zs=(Catch P0 P1,t)#cfg1" using Cons by fastforce
      have "(LanguageCon.com.Catch (fst x) P1, snd x) = lift_catch P1 x"
         by (simp add: lift_catch_def prod.case_eq_if)
      then have "LanguageCon.com.Catch (fst x) P1 = LanguageCon.com.Catch P0 P1 ∧ snd x = t"
         using Cons.prems(6) zs by fastforce           
      then show ?thesis
          by fastforce                
    qed
    then have step:"(n, Γ, (P0, t) # xs') ∈ cptn_mod_nest_call" using Cons by fastforce             
    have "fst (last ((P0, t) # xs')) = LanguageCon.com.Skip"
      using Cons.prems(3) x by auto
    then show ?case
      using Cons.prems(4) Cons.prems(6) CptnModNestCatch2.prems(1) 
            cptn_mod_nest_call.CptnModNestCatch2 local.step x by fastforce 
  qed          
next
  case (CptnModNestCatch3 n Γ P0 sa xs s' P1 ys zs) 
  thus ?case 
  proof (induct xs)
    case Nil thus ?case using Nil.prems(6) Nil.prems(7) by force
  next
    case (Cons x xs')
    then have x:"x=(P0,t)" 
    proof-
      have zs:"zs=(Catch P0 P1,t)#cfg1" using Cons by fastforce
      thus ?thesis using Cons(8) lift_catch_def unfolding lift_def
      proof -
        assume "zs = map (lift_catch P1) (x # xs') @ (P1, snd (last ((P0, Normal sa) # x # xs'))) # ys"
        then have "LanguageCon.com.Catch (fst x) P1 = LanguageCon.com.Catch P0 P1 ∧ snd x = t"
          by (simp add: case_prod_unfold lift_catch_def zs)          
        then show ?thesis
          by fastforce
      qed 
    qed
    then have step:"(n, Γ, (P0, t) # xs') ∈ cptn_mod_nest_call" using Cons by fastforce
    then obtain t' where t:"t=Normal t'" 
      using Normal_Normal Cons(2) Cons(5) cptn_mod_nest_cptn_mod cptn_eq_cptn_mod_set x
      by (metis snd_eqD)
    then show ?case 
    proof -
      have "last ((P0, Normal t') # xs') = last ((P0, Normal sa) # x # xs')"
        using t x by force
      then have "fst (last ((P0, Normal t') # xs')) = LanguageCon.com.Throw"
        using Cons.prems(3) by presburger
      then show ?thesis
        using Cons.prems(4) Cons.prems(5) Cons.prems(7) 
              CptnModNestCatch3.prems(1) cptn_mod_nest_call.CptnModNestCatch3 
              local.step t x by fastforce
    qed
  qed   
qed(fastforce+)


lemma elim_cptn_mod_nest_not_env_call:
 assumes a0:"(n,Γ,cfg) ∈  cptn_mod_nest_call" and
         a1:"cfg = (P,s)#(Q,t)#cfg1" and
         a2:"(∀f. redex P ≠ Call f) ∨  
             SmallStepCon.redex P = LanguageCon.com.Call fn ∧ Γ fn = None ∨
            (redex P = Call fn ∧ (∀sa. s≠Normal sa))"  
 shows "(n,Γ,(Q,t)#cfg1) ∈  cptn_mod_nest_call"
 using a0 a1 a2
proof (induct arbitrary: P Q cfg1 s t rule:cptn_mod_nest_call.induct )
case (CptnModNestSeq1 n Γ P0 s xs zs P1)
   then obtain P0' xs' where "xs =  (P0', t)#xs'" unfolding lift_def by fastforce
   then have step:"(n, Γ, (P0', t) # xs') ∈ cptn_mod_nest_call" using CptnModNestSeq1 by fastforce        
   have Q:"(Q, t) = lift P1 (P0', t) ∧ cfg1 = map (lift P1) xs'"
     using CptnModNestSeq1.hyps(3) CptnModNestSeq1.prems(1) ‹xs = (P0', t) # xs'› by auto
   also then have "(n, Γ, (LanguageCon.com.Seq P0' P1, t) # cfg1) ∈ cptn_mod_nest_call"
     by (meson cptn_mod_nest_call.CptnModNestSeq1 local.step)
   ultimately show ?case
     using CptnModNestSeq1.prems(1)
     by (simp add: Cons_lift Q)   
next
  case (CptnModNestSeq2 n Γ P0 sa xs P1 ys zs) 
  thus ?case 
  proof (induct xs)
    case Nil thus ?case using Nil.prems(6) Nil.prems(7) by force
  next
    case (Cons x xs')
    then have x:"∃P0'. x=(P0',t)" 
    proof-
      obtain P0'' where zs: "zs=(Seq P0'' P1,t)#cfg1" using Cons(7) Cons(8) 
        unfolding lift_def by (simp add: Cons_eq_append_conv case_prod_beta') 
      thus ?thesis using Cons(7) unfolding lift_def
      proof -
        assume "zs = map (λa. case a of (P, s) ⇒ (LanguageCon.com.Seq P P1, s)) (x # xs') @ 
                     (P1, snd (last ((P0, sa) # x # xs'))) # ys"
        then have "LanguageCon.com.Seq (fst x) P1 = LanguageCon.com.Seq P0'' P1 ∧ snd x = t"
          by (simp add: zs case_prod_beta)
        also have "sa=s" using Cons by fastforce
        ultimately show ?thesis by (meson eq_snd_iff)           
      qed 
    qed
    then obtain P0' where x:"x=(P0',t)" by auto
    then have step:"(n, Γ, (P0', t) # xs') ∈ cptn_mod_nest_call" using Cons by force
    have "fst (last ((P0', t) # xs')) = LanguageCon.com.Skip"
      using Cons.prems(3) x by force
    then show ?case
      using Cons.prems(4) Cons.prems(6) CptnModNestSeq2.prems(1) x 
           local.step cptn_mod_nest_call.CptnModNestSeq2[of n Γ P0' t xs' P1 ys] Cons_lift_append
           by (metis (no_types, lifting) last_ConsR list.inject list.simps(3))        
  qed          
next 
  case (CptnModNestSeq3 n Γ P0 sa xs s' ys zs P1) 
  thus ?case 
  proof (induct xs)
    case Nil thus ?case using Nil.prems(6) Nil.prems(7) by force
  next
    case (Cons x xs') 
    then have x:"∃P0'. x=(P0',t)"
    proof-
      obtain P0' where zs:"zs=(Seq P0' P1,t)#cfg1" using Cons(8) Cons(9) 
        unfolding lift_def
        unfolding lift_def by (simp add: Cons_eq_append_conv case_prod_beta')
      have "(LanguageCon.com.Seq (fst x) P1, snd x) = lift P1 x"
         by (simp add: lift_def prod.case_eq_if)
      then have "LanguageCon.com.Seq (fst x) P1 = LanguageCon.com.Seq P0' P1 ∧ snd x = t"
         using zs by (simp add: Cons.prems(7)) 
      then show ?thesis by (meson eq_snd_iff)                        
    qed
    then obtain P0' where x:"x=(P0',t)" by auto
    then have step:"(n, Γ, (P0', t) # xs') ∈ cptn_mod_nest_call"
    proof -
      have f1: "LanguageCon.com.Seq P0 P1 = P ∧ Normal sa = s"
        using CptnModNestSeq3.prems(1) by blast
      then have "SmallStepCon.redex P = SmallStepCon.redex P0"
        by (metis SmallStepCon.redex.simps(4))
      then show ?thesis
        using f1 Cons.prems(2) CptnModNestSeq3.prems(2) x by presburger
    qed      
    then obtain t' where t:"t=Normal t'" 
      using Normal_Normal Cons(2) Cons(5) cptn_mod_nest_cptn_mod cptn_eq_cptn_mod_set x
      by (metis snd_eqD)
    then show ?case using x Cons(5) Cons(6) cptn_mod_nest_call.CptnModNestSeq3 step
    proof -
      have "last ((P0', Normal t') # xs') = last ((P0, Normal sa) # x # xs')"
        using t x by force
      also then have "fst (last ((P0', Normal t') # xs')) = LanguageCon.com.Throw"
        using Cons.prems(3) by presburger
      ultimately show ?thesis
        using Cons.prems(4) Cons.prems(5) Cons.prems(7) 
              CptnModNestSeq3.prems(1) cptn_mod_nest_call.CptnModNestSeq3[of n Γ P0' t' xs' s' ys] 
              local.step t x  Cons_lift_append
      by (metis (no_types, lifting) list.sel(3))               
    qed       
  qed       
next
  case (CptnModNestCatch1 n Γ P0 s xs zs P1) 
   then obtain P0' xs' where xs:"xs =  (P0', t)#xs'" unfolding lift_catch_def by fastforce
   then have step:"(n, Γ, (P0', t) # xs') ∈ cptn_mod_nest_call" using CptnModNestCatch1 by fastforce        
   have Q:"(Q, t) = lift_catch P1 (P0', t) ∧ cfg1 = map (lift_catch P1) xs'"
      using CptnModNestCatch1.hyps(3) CptnModNestCatch1.prems(1) xs by auto
    then have "(n, Γ, (Catch P0' P1, t) # cfg1) ∈ cptn_mod_nest_call"
      by (meson cptn_mod_nest_call.CptnModNestCatch1 local.step)
    then show ?case
      using CptnModNestCatch1.prems(1) by (simp add:Cons_lift_catch Q)
next
  case (CptnModNestCatch2 n Γ P0 sa xs ys zs P1) 
  thus ?case 
  proof (induct xs)
    case Nil thus ?case using Nil.prems(6) Nil.prems(7) by force
  next
    case (Cons x xs') 
    then have x:"∃P0'. x=(P0',t)" 
    proof-
      obtain P0' where zs:"zs=(Catch P0' P1,t)#cfg1" using Cons unfolding lift_catch_def
        by (simp add: case_prod_unfold)
      have "(LanguageCon.com.Catch (fst x) P1, snd x) = lift_catch P1 x"
         by (simp add: lift_catch_def prod.case_eq_if)
      then have "LanguageCon.com.Catch (fst x) P1 = LanguageCon.com.Catch P0' P1 ∧ snd x = t"
         using Cons.prems(6) zs by fastforce           
      then show ?thesis by (meson eq_snd_iff)          
    qed
    then obtain P0' where x:"x=(P0',t)" by auto
    then have step:"(n, Γ, (P0', t) # xs') ∈ cptn_mod_nest_call"     
      using Cons.prems(2) CptnModNestCatch2.prems(1) CptnModNestCatch2.prems(2) x by force

    have skip:"fst (last ((P0', t) # xs')) = LanguageCon.com.Skip"
      using Cons.prems(3) x by auto
    show ?case
    proof -
      have "(P, s) # (Q, t) # cfg1 = (LanguageCon.com.Catch P0 P1, sa) # map (lift_catch P1) (x # xs') @ 
              (LanguageCon.com.Skip, snd (last ((P0, sa) # x # xs'))) # ys"
        using CptnModNestCatch2.prems  Cons.prems(6) by auto
      then show ?thesis 
        using Cons_lift_catch_append Cons.prems(4) 
              cptn_mod_nest_call.CptnModNestCatch2[OF local.step skip] last.simps list.distinct(1)
              x 
        by (metis (no_types)  list.sel(3) x)
    qed
  qed          
next
  case (CptnModNestCatch3 n Γ P0 sa xs s' P1 ys zs) 
  thus ?case 
  proof (induct xs)
    case Nil thus ?case using Nil.prems(6) Nil.prems(7) by force
  next
    case (Cons x xs')
    then have x:"∃P0'. x=(P0',t)" 
    proof-
      obtain P0' where zs:"zs=(Catch P0' P1,t)#cfg1" using Cons unfolding lift_catch_def
        by (simp add: case_prod_unfold)
      thus ?thesis using Cons(8) lift_catch_def unfolding lift_def
      proof -
        assume "zs = map (lift_catch P1) (x # xs') @ (P1, snd (last ((P0, Normal sa) # x # xs'))) # ys"
        then have "LanguageCon.com.Catch (fst x) P1 = LanguageCon.com.Catch P0' P1 ∧ snd x = t"
          by (simp add: case_prod_unfold lift_catch_def zs)          
        then show ?thesis by (meson eq_snd_iff)  
      qed 
    qed
    then obtain P0' where x:"x=(P0',t)" by auto
    then have step:"(n, Γ, (P0', t) # xs') ∈ cptn_mod_nest_call" using Cons
      using Cons.prems(2) CptnModNestCatch3.prems(1) CptnModNestCatch3.prems(2) x by force
    then obtain t' where t:"t=Normal t'" 
      using Normal_Normal Cons(2) Cons(5) cptn_mod_nest_cptn_mod cptn_eq_cptn_mod_set x
      by (metis snd_eqD)
    then show ?case 
    proof -
      have "last ((P0', Normal t') # xs') = last ((P0, Normal sa) # x # xs')"
        using t x by force
      also then have "fst (last ((P0', Normal t') # xs')) = LanguageCon.com.Throw"
        using Cons.prems(3) by presburger
      ultimately show ?thesis
        using Cons.prems(4) Cons.prems(5) Cons.prems(7) 
              CptnModNestCatch3.prems(1) cptn_mod_nest_call.CptnModNestCatch3[of n Γ P0' t' xs' s' P1] 
              local.step t x by (metis Cons_lift_catch_append list.sel(3)) 
    qed
  qed
next
case (CptnModNestWhile1 n Γ P0 s' xs b zs) 
  thus ?case
   using cptn_mod_nest_call.CptnModNestSeq1 list.inject by blast   
next
  case (CptnModNestWhile2 n Γ P0 s' xs b zs ys)  
  have "(LanguageCon.com.While b P0, Normal s') = (P, s) ∧ 
        (LanguageCon.com.Seq P0 (LanguageCon.com.While b P0), Normal s') # zs = (Q, t) # cfg1"
    using CptnModNestWhile2.prems by fastforce
  then show ?case
    using CptnModNestWhile2.hyps(1) CptnModNestWhile2.hyps(3) 
          CptnModNestWhile2.hyps(5) CptnModNestWhile2.hyps(6) 
          cptn_mod_nest_call.CptnModNestSeq2 by blast
next
  case (CptnModNestWhile3 n Γ P0 s' xs b zs) thus ?case
    by (metis (no_types) CptnModNestWhile3.hyps(1) CptnModNestWhile3.hyps(3) CptnModNestWhile3.hyps(5) 
                         CptnModNestWhile3.hyps(6) CptnModNestWhile3.hyps(8) CptnModNestWhile3.prems 
                         cptn_mod_nest_call.CptnModNestSeq3 list.inject)  
qed(fastforce+)

inductive_cases stepc_call_skip_normal:
 "Γ⊢c(Call p,Normal s) → (Skip,s')"

lemma elim_cptn_mod_nest_call_n_greater_zero:
 assumes a0:"(n,Γ,cfg) ∈  cptn_mod_nest_call" and
         a1:"cfg = (P,Normal s)#(Q,t)#cfg1 ∧ P = Call f ∧ Γ f = Some Q ∧ P≠Q"
 shows  "n>0" 
  using a0 a1 by (induct rule:cptn_mod_nest_call.induct, fastforce+)


lemma elim_cptn_mod_nest_call_0_False:
 assumes a0:"(0,Γ,cfg) ∈  cptn_mod_nest_call" and
         a1:"cfg = (P,Normal s)#(Q,t)#cfg1 ∧ P = Call f ∧ Γ f = Some Q ∧ P≠Q"
shows "PP"
using a0 a1 elim_cptn_mod_nest_call_n_greater_zero 
by fastforce

lemma elim_cptn_mod_nest_call_n_dec1:
 assumes a0:"(n,Γ,cfg) ∈  cptn_mod_nest_call" and
         a1:"cfg = (P,Normal s)#(Q,t)#cfg1 ∧ P = Call f ∧ Γ f = Some Q ∧ t= Normal s ∧ P≠Q"
 shows "(n-1,Γ,(Q,t)#cfg1) ∈  cptn_mod_nest_call"
 using a0 a1 
  by (induct rule:cptn_mod_nest_call.induct,fastforce+)

lemma elim_cptn_mod_nest_call_n_dec:
 assumes a0:"(n,Γ,(Call f,Normal s)#(the (Γ f),Normal s)#cfg1) ∈  cptn_mod_nest_call" and
         a1:"Γ f = Some Q " and a2:"Call f≠the (Γ f)"
       shows "(n-1,Γ,(the (Γ f),Normal s)#cfg1) ∈  cptn_mod_nest_call"
  using elim_cptn_mod_nest_call_n_dec1
 using a0 a1 a2
  by fastforce


lemma elim_cptn_mod_nest_call_n:
 assumes a0:"(n,Γ,cfg) ∈  cptn_mod_nest_call" and
         a1:"cfg = (P, s)#(Q,t)#cfg1"          
 shows "(n,Γ,(Q,t)#cfg1) ∈  cptn_mod_nest_call"
 using a0 a1 
proof (induct arbitrary: P Q cfg1 s t rule:cptn_mod_nest_call.induct )
case (CptnModNestCall n Γ bdy sa ys p)
  thus ?case using cptn_mod_nest_mono1 list.inject by blast
next 
case (CptnModNestSeq1 n Γ P0 s xs zs P1) 
   then obtain P0' xs' where "xs =  (P0', t)#xs'" unfolding lift_def by fastforce
   then have step:"(n, Γ, (P0', t) # xs') ∈ cptn_mod_nest_call" using CptnModNestSeq1 by fastforce        
   have Q:"(Q, t) = lift P1 (P0', t) ∧ cfg1 = map (lift P1) xs'"
     using CptnModNestSeq1.hyps(3) CptnModNestSeq1.prems(1) ‹xs = (P0', t) # xs'› by auto
   also then have "(n, Γ, (LanguageCon.com.Seq P0' P1, t) # cfg1) ∈ cptn_mod_nest_call"
     by (meson cptn_mod_nest_call.CptnModNestSeq1 local.step)
   ultimately show ?case
     using CptnModNestSeq1.prems(1)
     by (simp add: Cons_lift Q)   
next
  case (CptnModNestSeq2 n Γ P0 sa xs P1 ys zs) 
  thus ?case 
  proof (induct xs)
    case Nil thus ?case using Nil.prems(6) Nil.prems(7) by force
  next
    case (Cons x xs')
    then have x:"∃P0'. x=(P0',t)" 
    proof-
      obtain P0'' where zs: "zs=(Seq P0'' P1,t)#cfg1" using Cons(7) Cons(8) 
        unfolding lift_def by (simp add: Cons_eq_append_conv case_prod_beta') 
      thus ?thesis using Cons(7) unfolding lift_def
      proof -
        assume "zs = map (λa. case a of (P, s) ⇒ (LanguageCon.com.Seq P P1, s)) (x # xs') @ 
                     (P1, snd (last ((P0, sa) # x # xs'))) # ys"
        then have "LanguageCon.com.Seq (fst x) P1 = LanguageCon.com.Seq P0'' P1 ∧ snd x = t"
          by (simp add: zs case_prod_beta)
        also have "sa=s" using Cons by fastforce
        ultimately show ?thesis by (meson eq_snd_iff)           
      qed 
    qed
    then obtain P0' where x:"x=(P0',t)" by auto
    then have step:"(n, Γ, (P0', t) # xs') ∈ cptn_mod_nest_call" using Cons by force
    have "fst (last ((P0', t) # xs')) = LanguageCon.com.Skip"
      using Cons.prems(3) x by force
    then show ?case
      using Cons.prems(4) Cons.prems(6) CptnModNestSeq2.prems(1) x 
           local.step cptn_mod_nest_call.CptnModNestSeq2[of n Γ P0' t xs' P1 ys] Cons_lift_append
           by (metis (no_types, lifting) last_ConsR list.inject list.simps(3))        
  qed          
next 
  case (CptnModNestSeq3 n Γ P0 sa xs s' ys zs P1) 
  thus ?case 
  proof (induct xs)
    case Nil thus ?case using Nil.prems(6) Nil.prems(7) by force
  next
    case (Cons x xs') 
    then have x:"∃P0'. x=(P0',t)"
    proof-
      obtain P0' where zs:"zs=(Seq P0' P1,t)#cfg1" using Cons(8) Cons(9) 
        unfolding lift_def
        unfolding lift_def by (simp add: Cons_eq_append_conv case_prod_beta')
      have "(LanguageCon.com.Seq (fst x) P1, snd x) = lift P1 x"
         by (simp add: lift_def prod.case_eq_if)
      then have "LanguageCon.com.Seq (fst x) P1 = LanguageCon.com.Seq P0' P1 ∧ snd x = t"
         using zs by (simp add: Cons.prems(7)) 
      then show ?thesis by (meson eq_snd_iff)                        
    qed
    then obtain P0' where x:"x=(P0',t)" by auto
    then have step:"(n, Γ, (P0', t) # xs') ∈ cptn_mod_nest_call" using Cons by fastforce         
    then obtain t' where t:"t=Normal t'" 
      using Normal_Normal Cons(2) Cons(5) cptn_mod_nest_cptn_mod cptn_eq_cptn_mod_set x
      by (metis snd_eqD)
    then show ?case using x Cons(5) Cons(6) cptn_mod_nest_call.CptnModNestSeq3 step
    proof -
      have "last ((P0', Normal t') # xs') = last ((P0, Normal sa) # x # xs')"
        using t x by force
      also then have "fst (last ((P0', Normal t') # xs')) = LanguageCon.com.Throw"
        using Cons.prems(3) by presburger
      ultimately show ?thesis
        using Cons.prems(4) Cons.prems(5) Cons.prems(7) 
              CptnModNestSeq3.prems(1) cptn_mod_nest_call.CptnModNestSeq3[of n Γ P0' t' xs' s' ys] 
              local.step t x  Cons_lift_append
      by (metis (no_types, lifting) list.sel(3))               
    qed       
  qed       
next
  case (CptnModNestCatch1 n Γ P0 s xs zs P1) 
   then obtain P0' xs' where xs:"xs =  (P0', t)#xs'" unfolding lift_catch_def by fastforce
   then have step:"(n, Γ, (P0', t) # xs') ∈ cptn_mod_nest_call" using CptnModNestCatch1 by fastforce        
   have Q:"(Q, t) = lift_catch P1 (P0', t) ∧ cfg1 = map (lift_catch P1) xs'"
      using CptnModNestCatch1.hyps(3) CptnModNestCatch1.prems(1) xs by auto
    then have "(n, Γ, (Catch P0' P1, t) # cfg1) ∈ cptn_mod_nest_call"
      by (meson cptn_mod_nest_call.CptnModNestCatch1 local.step)
    then show ?case
      using CptnModNestCatch1.prems(1) by (simp add:Cons_lift_catch Q)
next
  case (CptnModNestCatch2 n Γ P0 sa xs ys zs P1) 
  thus ?case 
  proof (induct xs)
    case Nil thus ?case using Nil.prems(6) Nil.prems(7) by force
  next
    case (Cons x xs') 
    then have x:"∃P0'. x=(P0',t)" 
    proof-
      obtain P0' where zs:"zs=(Catch P0' P1,t)#cfg1" using Cons unfolding lift_catch_def
        by (simp add: case_prod_unfold)
      have "(LanguageCon.com.Catch (fst x) P1, snd x) = lift_catch P1 x"
         by (simp add: lift_catch_def prod.case_eq_if)
      then have "LanguageCon.com.Catch (fst x) P1 = LanguageCon.com.Catch P0' P1 ∧ snd x = t"
         using Cons.prems(6) zs by fastforce           
      then show ?thesis by (meson eq_snd_iff)          
    qed
    then obtain P0' where x:"x=(P0',t)" by auto
    then have step:"(n, Γ, (P0', t) # xs') ∈ cptn_mod_nest_call" using Cons by fastforce             
    have skip:"fst (last ((P0', t) # xs')) = LanguageCon.com.Skip"
      using Cons.prems(3) x by auto
    show ?case
    proof -
      have "(P, s) # (Q, t) # cfg1 = (LanguageCon.com.Catch P0 P1, sa) # map (lift_catch P1) (x # xs') @ 
              (LanguageCon.com.Skip, snd (last ((P0, sa) # x # xs'))) # ys"
        using CptnModNestCatch2.prems  Cons.prems(6) by auto
      then show ?thesis 
        using Cons_lift_catch_append Cons.prems(4) 
              cptn_mod_nest_call.CptnModNestCatch2[OF local.step skip] last.simps list.distinct(1)
              x 
        by (metis (no_types)  list.sel(3) x)
    qed
  qed          
next
  case (CptnModNestCatch3 n Γ P0 sa xs s' P1 ys zs) 
  thus ?case 
  proof (induct xs)
    case Nil thus ?case using Nil.prems(6) Nil.prems(7) by force
  next
    case (Cons x xs')
    then have x:"∃P0'. x=(P0',t)" 
    proof-
      obtain P0' where zs:"zs=(Catch P0' P1,t)#cfg1" using Cons unfolding lift_catch_def
        by (simp add: case_prod_unfold)
      thus ?thesis using Cons(8) lift_catch_def unfolding lift_def
      proof -
        assume "zs = map (lift_catch P1) (x # xs') @ (P1, snd (last ((P0, Normal sa) # x # xs'))) # ys"
        then have "LanguageCon.com.Catch (fst x) P1 = LanguageCon.com.Catch P0' P1 ∧ snd x = t"
          by (simp add: case_prod_unfold lift_catch_def zs)          
        then show ?thesis by (meson eq_snd_iff)  
      qed 
    qed
    then obtain P0' where x:"x=(P0',t)" by auto
    then have step:"(n, Γ, (P0', t) # xs') ∈ cptn_mod_nest_call" using Cons by fastforce
    then obtain t' where t:"t=Normal t'" 
      using Normal_Normal Cons(2) Cons(5) cptn_mod_nest_cptn_mod cptn_eq_cptn_mod_set x
      by (metis snd_eqD)
    then show ?case 
    proof -
      have "last ((P0', Normal t') # xs') = last ((P0, Normal sa) # x # xs')"
        using t x by force
      also then have "fst (last ((P0', Normal t') # xs')) = LanguageCon.com.Throw"
        using Cons.prems(3) by presburger
      ultimately show ?thesis
        using Cons.prems(4) Cons.prems(5) Cons.prems(7) 
              CptnModNestCatch3.prems(1) cptn_mod_nest_call.CptnModNestCatch3[of n Γ P0' t' xs' s' P1] 
              local.step t x by (metis Cons_lift_catch_append list.sel(3)) 
    qed
  qed
next
case (CptnModNestWhile1 n Γ P0 s' xs b zs) 
  thus ?case
   using cptn_mod_nest_call.CptnModNestSeq1 list.inject by blast   
next
  case (CptnModNestWhile2 n Γ P0 s' xs b zs ys)  
  have "(LanguageCon.com.While b P0, Normal s') = (P, s) ∧ 
        (LanguageCon.com.Seq P0 (LanguageCon.com.While b P0), Normal s') # zs = (Q, t) # cfg1"
    using CptnModNestWhile2.prems by fastforce
  then show ?case
    using CptnModNestWhile2.hyps(1) CptnModNestWhile2.hyps(3) 
          CptnModNestWhile2.hyps(5) CptnModNestWhile2.hyps(6) 
          cptn_mod_nest_call.CptnModNestSeq2 by blast
next
  case (CptnModNestWhile3 n Γ P0 s' xs b zs) thus ?case
    by (metis (no_types) CptnModNestWhile3.hyps(1) CptnModNestWhile3.hyps(3) CptnModNestWhile3.hyps(5) 
                         CptnModNestWhile3.hyps(6) CptnModNestWhile3.hyps(8) CptnModNestWhile3.prems 
                         cptn_mod_nest_call.CptnModNestSeq3 list.inject)  
qed (fastforce+) 




definition min_call where
"min_call n Γ cfs ≡ (n,Γ,cfs) ∈  cptn_mod_nest_call ∧ (∀m<n. ¬((m,Γ,cfs) ∈  cptn_mod_nest_call))"

lemma minimum_nest_call:
  "(m,Γ,cfs) ∈ cptn_mod_nest_call ⟹
   ∃n. min_call n Γ cfs"
unfolding min_call_def
proof (induct arbitrary: m rule:cptn_mod_nest_call.induct) 
 case (CptnModNestOne) thus ?case using cptn_mod_nest_call.CptnModNestOne by blast 
next
  case (CptnModNestEnv Γ P s t n xs) 
  then have "¬ Γ⊢c (P, s) → (P, t)" 
   using mod_env_not_component step_change_p_or_eq_s by blast 
  then obtain min_n where min:"(min_n, Γ, (P, t) # xs) ∈ cptn_mod_nest_call ∧ 
                             (∀m<min_n.  (m, Γ, (P, t) # xs) ∉ cptn_mod_nest_call)" 
    using CptnModNestEnv by blast
  then have  "(min_n, Γ, (P,s)#(P, t) # xs) ∈ cptn_mod_nest_call"     
    using cptn_mod_nest_call.CptnModNestEnv CptnModNestEnv by blast
  also have "(∀m<min_n. (m, Γ, (P, s)#(P, t) # xs) ∉ cptn_mod_nest_call)"
    using elim_cptn_mod_nest_call_n min by fastforce    
  ultimately show ?case by auto  
next
  case (CptnModNestSkip Γ P s t n xs)    
  then obtain min_n where 
     min:"(min_n, Γ, (LanguageCon.com.Skip, t) # xs) ∈ cptn_mod_nest_call ∧
        (∀m<min_n. (m, Γ, (LanguageCon.com.Skip, t) # xs) ∉ cptn_mod_nest_call)" 
    by auto
  then have "(min_n, Γ, (P,s)#(LanguageCon.com.Skip, t) # xs) ∈ cptn_mod_nest_call"
    using cptn_mod_nest_call.CptnModNestSkip CptnModNestSkip by blast
  also have "(∀m<min_n. (m, Γ, (P, s)#(LanguageCon.com.Skip, t) # xs) ∉ cptn_mod_nest_call)"
    using elim_cptn_mod_nest_call_n min by blast      
  ultimately show ?case by fastforce   
next
   case (CptnModNestThrow Γ P s t n xs) thus ?case
     by (meson cptn_mod_nest_call.CptnModNestThrow elim_cptn_mod_nest_call_n)    
next
  case (CptnModNestCondT n Γ P0 s xs b P1) thus ?case
    by (meson cptn_mod_nest_call.CptnModNestCondT elim_cptn_mod_nest_call_n)
next
  case (CptnModNestCondF n Γ P1 s xs b P0) thus ?case
    by (meson cptn_mod_nest_call.CptnModNestCondF elim_cptn_mod_nest_call_n)
next
  case (CptnModNestSeq1 n Γ P s xs zs Q) thus ?case
    by (metis (no_types, lifting) Seq_P_Not_finish cptn_mod_nest_call.CptnModNestSeq1 div_seq_nest)
next
  case (CptnModNestSeq2 n Γ P s xs Q ys zs) 
  then obtain min_p where 
     min_p:"(min_p, Γ, (P,  s) # xs) ∈ cptn_mod_nest_call ∧
        (∀m<min_p. (m, Γ, (P, s) # xs) ∉ cptn_mod_nest_call)" 
    by auto 
  from CptnModNestSeq2(5) obtain min_q where
    min_q:"(min_q, Γ, (Q, snd (last ((P, s) # xs))) # ys) ∈ cptn_mod_nest_call ∧
        (∀m<min_q. (m, Γ, (Q, snd (last ((P, s) # xs))) # ys) ∉ cptn_mod_nest_call)"
  by auto
  thus ?case
  proof(cases "min_p≥min_q")
    case True
    then have "(min_p, Γ, (Q, snd (last ((P,s) # xs))) # ys) ∈ cptn_mod_nest_call"
      using min_q using cptn_mod_nest_mono by blast 
    then have "(min_p, Γ, (Seq P Q, s) # zs) ∈ cptn_mod_nest_call"
      using conjunct1[OF min_p] cptn_mod_nest_call.CptnModNestSeq2[of min_p Γ P s xs Q ys zs] 
            CptnModNestSeq2(6)  CptnModNestSeq2(3)
    by blast
    also have "∀m<min_p. (m, Γ,(Seq P Q,s) # zs) ∉ cptn_mod_nest_call"
      by (metis CptnModNestSeq2.hyps(3) CptnModNestSeq2.hyps(6) Seq_P_Ends_Normal div_seq_nest min_p)      
    ultimately show ?thesis by auto
  next
    case False 
    then have "(min_q, Γ, (P,  s) # xs) ∈ cptn_mod_nest_call"
      using min_p cptn_mod_nest_mono by force 
    then have "(min_q, Γ, (Seq P Q, s) # zs) ∈ cptn_mod_nest_call"
      using conjunct1[OF min_q] cptn_mod_nest_call.CptnModNestSeq2[of min_q Γ P s xs Q ys zs] 
            CptnModNestSeq2(6)  CptnModNestSeq2(3)
    by blast
    also have "∀m<min_q. (m, Γ,(Seq P Q,s) # zs) ∉ cptn_mod_nest_call"
     proof -
      {fix m
      assume min_m:"m<min_q"
      then have "(m, Γ,(Seq P Q, s) # zs) ∉ cptn_mod_nest_call"
      proof -
      {assume ass:"(m, Γ, (Seq P Q, s) # zs) ∈ cptn_mod_nest_call"
       then obtain xs' s' s'' where 
          m_cptn:"(m, Γ, (P, s) # xs') ∈ cptn_mod_nest_call ∧ 
                   seq_cond_nest zs Q xs' P s s'' s' Γ m" 
         using  
          div_seq_nest[of m Γ "(LanguageCon.com.Seq P Q, s) # zs"]
          by fastforce
       then have "seq_cond_nest zs Q xs' P s s'' s' Γ m" by auto
       then have ?thesis
         using Seq_P_Ends_Normal[OF CptnModNestSeq2(6) CptnModNestSeq2(3) ass]
               min_m min_q 
         by (metis last_length)          
      } thus ?thesis by auto
      qed
      }thus ?thesis by auto
    qed  
    ultimately show ?thesis by auto
  qed
next
  case (CptnModNestSeq3 n Γ P s xs s' ys zs Q) 
  then obtain min_p where 
     min_p:"(min_p, Γ, (P, Normal s) # xs) ∈ cptn_mod_nest_call ∧
        (∀m<min_p. (m, Γ, (P, Normal s) # xs) ∉ cptn_mod_nest_call)" 
    by auto 
  from CptnModNestSeq3(6) obtain min_q where
    min_q:"(min_q, Γ, (Throw, Normal s') # ys) ∈ cptn_mod_nest_call ∧
        (∀m<min_q. (m, Γ, (Throw, Normal s') # ys) ∉ cptn_mod_nest_call)"
  by auto
  thus ?case
  proof(cases "min_p≥min_q")
    case True 
    then have "(min_p, Γ, (Throw, Normal s') # ys) ∈ cptn_mod_nest_call"
      using min_q using cptn_mod_nest_mono by blast 
    then have "(min_p, Γ, (Seq P Q, Normal s) # zs) ∈ cptn_mod_nest_call"
      using conjunct1[OF min_p] cptn_mod_nest_call.CptnModNestSeq3[of min_p Γ P s xs s' ys zs Q] 
            CptnModNestSeq3(4)  CptnModNestSeq3(3) CptnModNestSeq3(7)
    by blast
    also have "∀m<min_p. (m, Γ,(Seq P Q,Normal s) # zs) ∉ cptn_mod_nest_call"
      by (metis CptnModNestSeq3.hyps(3) CptnModNestSeq3.hyps(4) CptnModNestSeq3.hyps(7) Seq_P_Ends_Abort div_seq_nest min_p)
    ultimately show ?thesis by auto
  next
    case False
    then have "(min_q, Γ, (P,  Normal s) # xs) ∈ cptn_mod_nest_call"
      using min_p cptn_mod_nest_mono by force 
    then have "(min_q, Γ, (Seq P Q, Normal s) # zs) ∈ cptn_mod_nest_call"
      using conjunct1[OF min_q] cptn_mod_nest_call.CptnModNestSeq3[of min_q Γ P s xs s' ys zs Q] 
            CptnModNestSeq3(4)  CptnModNestSeq3(3) CptnModNestSeq3(7)
    by blast
    also have "∀m<min_q. (m, Γ,(Seq P Q,Normal s) # zs) ∉ cptn_mod_nest_call"
      by (metis CptnModNestSeq3.hyps(3) CptnModNestSeq3.hyps(4) CptnModNestSeq3.hyps(7) Seq_P_Ends_Abort div_seq_nest min_q)     
    ultimately show ?thesis by auto
  qed 
next
  case (CptnModNestWhile1 n Γ P s xs b zs) 
  then obtain min_n where 
     min:"(min_n, Γ, (P, Normal s) # xs) ∈ cptn_mod_nest_call ∧
        (∀m<min_n. (m, Γ, (P, Normal s) # xs) ∉ cptn_mod_nest_call)" 
    by auto
  then have "(min_n, Γ, (While b P, Normal s) # (Seq P (While b P), Normal s) # zs) ∈ cptn_mod_nest_call"
    using cptn_mod_nest_call.CptnModNestWhile1[of min_n Γ P s xs b zs] CptnModNestWhile1
    by meson 
  also have "∀m<min_n. (m, Γ,(While b P, Normal s) # (Seq P (While b P), Normal s) # zs) ∉ cptn_mod_nest_call"
    by (metis CptnModNestWhile1.hyps(4) Seq_P_Not_finish div_seq_nest elim_cptn_mod_nest_call_n min) 
  ultimately show ?case by auto
next
  case (CptnModNestWhile2 n Γ P s xs b zs ys) 
  then obtain min_n_p where 
     min_p:"(min_n_p, Γ, (P, Normal s) # xs) ∈ cptn_mod_nest_call ∧
        (∀m<min_n_p. (m, Γ, (P, Normal s) # xs) ∉ cptn_mod_nest_call)" 
    by auto
  from CptnModNestWhile2 obtain min_n_w where
     min_w:"(min_n_w, Γ, (LanguageCon.com.While b P, snd (last ((P, Normal s) # xs))) # ys) ∈ cptn_mod_nest_call ∧
        (∀m<min_n_w. (m, Γ, (LanguageCon.com.While b P, snd (last ((P, Normal s) # xs))) # ys)
               ∉ cptn_mod_nest_call)"
    by auto
  thus ?case 
  proof (cases "min_n_p≥min_n_w")
    case True 
    then have "(min_n_p, Γ, 
      (LanguageCon.com.While b P, snd (last ((P, Normal s) # xs))) # ys) ∈ cptn_mod_nest_call"
      using min_w using cptn_mod_nest_mono by blast 
    then have "(min_n_p, Γ, (While b P, Normal s) # (Seq P (While b P), Normal s) # zs) ∈ cptn_mod_nest_call"
      using min_p cptn_mod_nest_call.CptnModNestWhile2[of min_n_p Γ P s xs b zs] CptnModNestWhile2
      by blast
    also have "∀m<min_n_p. (m, Γ,(While b P, Normal s) # (Seq P (While b P), Normal s) # zs) ∉ cptn_mod_nest_call"
      by (metis CptnModNestWhile2.hyps(3) CptnModNestWhile2.hyps(5) 
                Seq_P_Ends_Normal div_seq_nest elim_cptn_mod_nest_call_n min_p)    
    ultimately show ?thesis by auto  
  next
    case False
    then have False:"min_n_p<min_n_w" by auto
    then have "(min_n_w, Γ, (P, Normal s) # xs) ∈ cptn_mod_nest_call"
      using min_p  cptn_mod_nest_mono by force 
    then have "(min_n_w, Γ, (While b P, Normal s) # (Seq P (While b P), Normal s) # zs) ∈ cptn_mod_nest_call"
      using min_w min_p cptn_mod_nest_call.CptnModNestWhile2[of min_n_w Γ P s xs b zs] CptnModNestWhile2
      by blast
    also have "∀m<min_n_w. (m, Γ,(While b P, Normal s) # (Seq P (While b P), Normal s) # zs) ∉ cptn_mod_nest_call"      
    proof -
      {fix m
      assume min_m:"m<min_n_w"
      then have "(m, Γ,(While b P, Normal s) # (Seq P (While b P), Normal s) # zs) ∉ cptn_mod_nest_call"
      proof -
      {assume "(m, Γ,(While b P, Normal s) # (Seq P (While b P), Normal s) # zs) ∈ cptn_mod_nest_call"
       then have a1:"(m, Γ,(Seq P (While b P), Normal s) # zs) ∈ cptn_mod_nest_call"      
         using elim_cptn_mod_nest_not_env_call by fastforce
       then obtain xs' s' s'' where 
          m_cptn:"(m, Γ, (P, Normal s) # xs') ∈ cptn_mod_nest_call  ∧ 
                   seq_cond_nest zs (While b P) xs' P (Normal s) s'' s' Γ m" 
         using  
          div_seq_nest[of m Γ "(LanguageCon.com.Seq P (LanguageCon.com.While b P), Normal s) # zs"]
          by fastforce
       then have "seq_cond_nest zs (While b P) xs' P (Normal s) s'' s' Γ m" by auto
       then have ?thesis unfolding seq_cond_nest_def
         by (metis CptnModNestWhile2.hyps(3) CptnModNestWhile2.hyps(5) Seq_P_Ends_Normal a1 last_length m_cptn min_m min_w)           
     } thus ?thesis by auto
     qed
     }thus ?thesis by auto
    qed
    ultimately show ?thesis by auto
  qed
next
  case (CptnModNestWhile3 n Γ P s xs b s' ys zs) 
  then obtain min_n_p where 
     min_p:"(min_n_p, Γ, (P, Normal s) # xs) ∈ cptn_mod_nest_call ∧
        (∀m<min_n_p. (m, Γ, (P, Normal s) # xs) ∉ cptn_mod_nest_call)" 
    by auto
  from CptnModNestWhile3 obtain min_n_w where
     min_w:"(min_n_w, Γ, (Throw, snd (last ((P, Normal s) # xs))) # ys) ∈ cptn_mod_nest_call ∧
        (∀m<min_n_w. (m, Γ, (Throw, snd (last ((P, Normal s) # xs))) # ys)
               ∉ cptn_mod_nest_call)"
    by auto
  thus ?case 
  proof (cases "min_n_p≥min_n_w")
    case True 
    then have "(min_n_p, Γ, 
      (Throw, snd (last ((P, Normal s) # xs))) # ys) ∈ cptn_mod_nest_call"
      using min_w using cptn_mod_nest_mono by blast 
    then have "(min_n_p, Γ, (While b P, Normal s) # (Seq P (While b P), Normal s) # zs) ∈ cptn_mod_nest_call"
      using min_p cptn_mod_nest_call.CptnModNestWhile3[of min_n_p Γ P s xs b s' ys zs] 
            CptnModNestWhile3
      by fastforce
    also have "∀m<min_n_p. (m, Γ,(While b P, Normal s) # (Seq P (While b P), Normal s) # zs) ∉ cptn_mod_nest_call"
      by (metis CptnModNestWhile3.hyps(3) CptnModNestWhile3.hyps(5) CptnModNestWhile3.hyps(8) 
            Seq_P_Ends_Abort div_seq_nest elim_cptn_mod_nest_call_n min_p)    
    ultimately show ?thesis by auto  
  next
    case False
    then have False:"min_n_p<min_n_w" by auto
    then have "(min_n_w, Γ, (P, Normal s) # xs) ∈ cptn_mod_nest_call"
      using min_p  cptn_mod_nest_mono by force 
    then have "(min_n_w, Γ, (While b P, Normal s) # (Seq P (While b P), Normal s) # zs) ∈ cptn_mod_nest_call"
      using min_w min_p cptn_mod_nest_call.CptnModNestWhile3[of min_n_w Γ P s xs b s' ys zs] 
            CptnModNestWhile3
      by fastforce      
    also have "∀m<min_n_w. (m, Γ,(While b P, Normal s) # (Seq P (While b P), Normal s) # zs) ∉ cptn_mod_nest_call"
    proof -
      {fix m
      assume min_m:"m<min_n_w"
      then have "(m, Γ,(While b P, Normal s) # (Seq P (While b P), Normal s) # zs) ∉ cptn_mod_nest_call"
      proof -
      {assume "(m, Γ,(While b P, Normal s) # (Seq P (While b P), Normal s) # zs) ∈ cptn_mod_nest_call"
       then have s1:"(m, Γ,(Seq P (While b P), Normal s) # zs) ∈ cptn_mod_nest_call"      
         using elim_cptn_mod_nest_not_env_call by fastforce
       then obtain xs' s' s'' where 
          m_cptn:"(m, Γ, (P, Normal s) # xs') ∈ cptn_mod_nest_call  ∧ 
                   seq_cond_nest zs (While b P) xs' P (Normal s) s'' s' Γ m" 
         using  
          div_seq_nest[of m Γ "(LanguageCon.com.Seq P (LanguageCon.com.While b P), Normal s) # zs"]
          by fastforce
       then have "seq_cond_nest zs (While b P) xs' P (Normal s) s'' s' Γ m" by auto
       then have ?thesis unfolding seq_cond_nest_def
         by (metis CptnModNestWhile3.hyps(3) CptnModNestWhile3.hyps(5) CptnModNestWhile3.hyps(8) Seq_P_Ends_Abort s1 m_cptn min_m min_w)       
     } thus ?thesis by auto
     qed
     }thus ?thesis by auto
    qed
    ultimately show ?thesis by auto
  qed
next
  case (CptnModNestCall n Γ bdy s xs f) thus ?case
  proof -
    { fix nn :: "nat ⇒ nat"
     obtain nna :: nat where
      ff1: "(nna, Γ, (bdy, Normal s) # xs) ∈ cptn_mod_nest_call ∧ (∀n. ¬ n < nna ∨ (n, Γ, (bdy, Normal s) # xs) ∉ cptn_mod_nest_call)"
      by (meson CptnModNestCall.hyps(2))
    moreover
    { assume "(nn (nn (Suc nna)), Γ, (bdy, Normal s) # xs) ∈ cptn_mod_nest_call"
      then have "¬ Suc (nn (nn (Suc nna))) < Suc nna"
        using ff1 by blast
      then have "(nn (Suc nna), Γ, (LanguageCon.com.Call f, Normal s) # (bdy, Normal s) # xs) ∈ cptn_mod_nest_call ⟶ (∃n. (n, Γ, (LanguageCon.com.Call f, Normal s) # (bdy, Normal s) # xs) ∈ cptn_mod_nest_call ∧ 
                (¬ nn n < n ∨ (nn n, Γ, (LanguageCon.com.Call f, Normal s) # (bdy, Normal s) # xs) ∉ cptn_mod_nest_call))"
        using ff1 by (meson CptnModNestCall.hyps(3) CptnModNestCall.hyps(4) cptn_mod_nest_call.CptnModNestCall less_trans_Suc) }
    ultimately have "∃n. (n, Γ, (LanguageCon.com.Call f, Normal s) # (bdy, Normal s) # xs) ∈ cptn_mod_nest_call ∧ (¬ nn n < n ∨ (nn n, Γ, (LanguageCon.com.Call f, Normal s) # (bdy, Normal s) # xs) ∉ cptn_mod_nest_call)"
      by (metis (no_types) CptnModNestCall.hyps(3) CptnModNestCall.hyps(4) cptn_mod_nest_call.CptnModNestCall elim_cptn_mod_nest_call_n) }
   then show ?thesis
     by meson
  qed     
next 
 case (CptnModNestDynCom n Γ c s xs) thus ?case
   by (meson cptn_mod_nest_call.CptnModNestDynCom elim_cptn_mod_nest_call_n)
next
  case (CptnModNestGuard n Γ c s xs g f) thus ?case 
    by (meson cptn_mod_nest_call.CptnModNestGuard elim_cptn_mod_nest_call_n)   
next
 case (CptnModNestCatch1 n Γ P s xs  zs Q) thus ?case
   by (metis (no_types, lifting) Catch_P_Not_finish cptn_mod_nest_call.CptnModNestCatch1 div_catch_nest)
next
 case (CptnModNestCatch2 n Γ P s xs ys zs Q) 
 then obtain min_p where 
     min_p:"(min_p, Γ, (P,  s) # xs) ∈ cptn_mod_nest_call ∧
        (∀m<min_p. (m, Γ, (P, s) # xs) ∉ cptn_mod_nest_call)" 
    by auto 
  from CptnModNestCatch2(5) obtain min_q where
    min_q:"(min_q, Γ, (Skip, snd (last ((P, s) # xs))) # ys) ∈ cptn_mod_nest_call ∧
        (∀m<min_q. (m, Γ, (Skip, snd (last ((P, s) # xs))) # ys) ∉ cptn_mod_nest_call)"
  by auto
  thus ?case
  proof(cases "min_p≥min_q")
    case True
    then have "(min_p, Γ, (Skip, snd (last ((P,s) # xs))) # ys) ∈ cptn_mod_nest_call"
      using min_q using cptn_mod_nest_mono by blast 
    then have "(min_p, Γ, (Catch P Q, s) # zs) ∈ cptn_mod_nest_call"
      using conjunct1[OF min_p] cptn_mod_nest_call.CptnModNestCatch2[of min_p Γ P s xs] 
            CptnModNestCatch2(6)  CptnModNestCatch2(3)
    by blast
    also have "∀m<min_p. (m, Γ,(Catch P Q,s) # zs) ∉ cptn_mod_nest_call"
     proof -
      {fix m
      assume min_m:"m<min_p"
      then have "(m, Γ,(Catch P Q, s) # zs) ∉ cptn_mod_nest_call"
      proof -
      {assume ass:"(m, Γ, (Catch P Q, s) # zs) ∈ cptn_mod_nest_call"
       then obtain xs' s' s'' where 
          m_cptn:"(m, Γ, (P, s) # xs') ∈ cptn_mod_nest_call ∧ 
                   catch_cond_nest zs Q xs' P s s'' s' Γ m" 
         using  
          div_catch_nest[of m Γ "(Catch P Q, s) # zs"]
          by fastforce
       then have "catch_cond_nest zs Q xs' P s s'' s' Γ m" by auto
       then have "xs=xs'" 
         using Catch_P_Ends_Skip[OF CptnModNestCatch2(6) CptnModNestCatch2(3)]   
         by fastforce
       then have "(m, Γ, (P,s) # xs) ∈ cptn_mod_nest_call"
         using m_cptn by auto
       then have False using min_p min_m by fastforce
    } thus ?thesis by auto
    qed
    }thus ?thesis by auto
  qed
  ultimately show ?thesis by auto
  next
    case False
    then have "(min_q, Γ, (P,  s) # xs) ∈ cptn_mod_nest_call"
      using min_p cptn_mod_nest_mono by force 
    then have "(min_q, Γ, (Catch P Q, s) # zs) ∈ cptn_mod_nest_call"
      using conjunct1[OF min_q] cptn_mod_nest_call.CptnModNestCatch2[of min_q Γ P s xs ] 
            CptnModNestCatch2(6)  CptnModNestCatch2(3)
    by blast
    also have "∀m<min_q. (m, Γ,(Catch P Q,s) # zs) ∉ cptn_mod_nest_call"
     proof -
      {fix m
      assume min_m:"m<min_q"
      then have "(m, Γ,(Catch P Q, s) # zs) ∉ cptn_mod_nest_call"
      proof -
      {assume ass:"(m, Γ, (Catch P Q, s) # zs) ∈ cptn_mod_nest_call"
       then obtain xs' s' s'' where 
          m_cptn:"(m, Γ, (P, s) # xs') ∈ cptn_mod_nest_call ∧ 
                   catch_cond_nest zs Q xs' P s s'' s' Γ m" 
         using  
          div_catch_nest[of m Γ "(Catch P Q, s) # zs"]
          by fastforce
       then have "catch_cond_nest zs Q xs' P s s'' s' Γ m" by auto
       then have ?thesis
         using Catch_P_Ends_Skip[OF CptnModNestCatch2(6) CptnModNestCatch2(3)]
               min_m min_q 
       by blast         
      } thus ?thesis by auto
      qed
      }thus ?thesis by auto
    qed  
    ultimately show ?thesis by auto
  qed
next
 case (CptnModNestCatch3 n Γ P s xs s' Q ys zs ) then obtain min_p where 
     min_p:"(min_p, Γ, (P, Normal s) # xs) ∈ cptn_mod_nest_call ∧
        (∀m<min_p. (m, Γ, (P, Normal s) # xs) ∉ cptn_mod_nest_call)" 
    by auto 
  from CptnModNestCatch3(6)  CptnModNestCatch3(4) obtain min_q where
    min_q:"(min_q, Γ, (Q,  snd (last ((P, Normal s) # xs))) # ys) ∈ cptn_mod_nest_call ∧
        (∀m<min_q. (m, Γ, (Q,  snd (last ((P, Normal s) # xs))) # ys) ∉ cptn_mod_nest_call)"
  by auto
  thus ?case
  proof(cases "min_p≥min_q")
    case True
    then have "(min_p, Γ, (Q,  snd (last ((P, Normal s) # xs))) # ys) ∈ cptn_mod_nest_call"
      using min_q using cptn_mod_nest_mono by blast 
    then have "(min_p, Γ, (Catch P Q, Normal s) # zs) ∈ cptn_mod_nest_call"
      using conjunct1[OF min_p] cptn_mod_nest_call.CptnModNestCatch3[of min_p Γ P s xs s' Q ys zs] 
            CptnModNestCatch3(4)  CptnModNestCatch3(3) CptnModNestCatch3(7)
    by fastforce
    also have "∀m<min_p. (m, Γ,(Catch P Q,Normal s) # zs) ∉ cptn_mod_nest_call"
     proof -
      {fix m
      assume min_m:"m<min_p"
      then have "(m, Γ,(Catch P Q, Normal s) # zs) ∉ cptn_mod_nest_call"
      proof -
      {assume ass:"(m, Γ, (Catch P Q,Normal s) # zs) ∈ cptn_mod_nest_call"
       then obtain xs' ns' ns'' where 
          m_cptn:"(m, Γ, (P, Normal s) # xs') ∈ cptn_mod_nest_call ∧ 
                   catch_cond_nest zs Q xs' P (Normal s) ns'' ns' Γ m" 
         using  
          div_catch_nest[of m Γ "(Catch P Q, Normal s) # zs"]
          by fastforce
       then have "catch_cond_nest zs Q xs' P (Normal s) ns'' ns' Γ m" by auto       
       then have "xs=xs'" 
         using  Catch_P_Ends_Normal[OF CptnModNestCatch3(7)  CptnModNestCatch3(3) CptnModNestCatch3(4)]   
         by fastforce
       then have "(m, Γ, (P,Normal s) # xs) ∈ cptn_mod_nest_call"
         using m_cptn by auto
       then have False using min_p min_m by fastforce
    } thus ?thesis by auto
    qed
    }thus ?thesis by auto
  qed
  ultimately show ?thesis by auto
  next
    case False
    then have "(min_q, Γ, (P,  Normal s) # xs) ∈ cptn_mod_nest_call"
      using min_p cptn_mod_nest_mono by force 
    then have "(min_q, Γ, (Catch P Q, Normal s) # zs) ∈ cptn_mod_nest_call"
      using conjunct1[OF min_q] cptn_mod_nest_call.CptnModNestCatch3[of min_q Γ P s xs s' ] 
            CptnModNestCatch3(4)  CptnModNestCatch3(3) CptnModNestCatch3(7)
    by blast
    also have "∀m<min_q. (m, Γ,(Catch P Q,Normal s) # zs) ∉ cptn_mod_nest_call"
     proof -
      {fix m
      assume min_m:"m<min_q"
      then have "(m, Γ,(Catch P Q, Normal s) # zs) ∉ cptn_mod_nest_call"
      proof -
      {assume ass:"(m, Γ, (Catch P Q, Normal s) # zs) ∈ cptn_mod_nest_call"
       then obtain xs' ns' ns'' where 
          m_cptn:"(m, Γ, (P, Normal s) # xs') ∈ cptn_mod_nest_call ∧ 
                   catch_cond_nest zs Q xs' P (Normal s) ns'' ns' Γ m" 
         using  
          div_catch_nest[of m Γ "(Catch P Q, Normal s) # zs"]
          by fastforce
       then have "catch_cond_nest zs Q xs' P (Normal s) ns'' ns' Γ m" by auto
       then have ?thesis
         using Catch_P_Ends_Normal[OF CptnModNestCatch3(7) CptnModNestCatch3(3)  CptnModNestCatch3(4)]
               min_m min_q 
         by (metis last_length)                                
      } thus ?thesis by auto
      qed
      }thus ?thesis by auto
    qed  
    ultimately show ?thesis by auto
  qed  
qed

  lemma elim_cptn_mod_min_nest_call:
 assumes a0:"min_call n Γ cfg" and
         a1:"cfg = (P,s)#(Q,t)#cfg1" and
         a2:"(∀f. redex P ≠ Call f) ∨  
             SmallStepCon.redex P = LanguageCon.com.Call fn ∧ Γ fn = None ∨
            (redex P = Call fn ∧ (∀sa. s≠Normal sa)) ∨
            (redex P = Call fn ∧ P=Q)"     
 shows "min_call n Γ ((Q,t)#cfg1)"
proof -
  have a0: "(n,Γ,cfg) ∈  cptn_mod_nest_call" and
       a0': "(∀m<n. (m, Γ, cfg) ∉ cptn_mod_nest_call)"
  using a0 unfolding min_call_def by auto
  then have "(n,Γ,(Q,t)#cfg1) ∈  cptn_mod_nest_call"  
    using a0 a1 elim_cptn_mod_nest_call_n by blast
  also have "(∀m<n. (m, Γ, (Q,t)#cfg1) ∉ cptn_mod_nest_call)"      
  proof-
  { assume "¬(∀m<n. (m, Γ, (Q,t)#cfg1) ∉ cptn_mod_nest_call)"
    then obtain m where 
      asm0:"m<n" and 
      asm1:"(m, Γ, (Q,t)#cfg1) ∈ cptn_mod_nest_call"
    by auto
    then have "(m, Γ, cfg) ∈ cptn_mod_nest_call" 
      using  a0 a1 a2  cptn_mod_nest_cptn_mod cptn_if_cptn_mod cptn_mod_nest_call.CptnModNestEnv
          cptn_elim_cases(2)  not_func_redex_cptn_mod_nest_n'          
      by (metis (no_types, lifting) mod_env_not_component)

    then have False using a0' asm0 by auto
  } thus ?thesis by auto qed
  ultimately show ?thesis unfolding min_call_def by auto
qed

  

lemma elim_call_cptn_mod_min_nest_call:
 assumes a0:"min_call n Γ cfg" and
         a1:"cfg = (P,s)#(Q,t)#cfg1" and
         a2:"P = Call f ∧  
             Γ f = Some Q ∧ (∃sa. s=Normal sa) ∧ P≠Q"          
 shows "min_call (n-1) Γ ((Q,t)#cfg1)"
proof -
  obtain s' where a0: "(n,Γ,cfg) ∈  cptn_mod_nest_call" and
       a0': "(∀m<n. (m, Γ, cfg) ∉ cptn_mod_nest_call)" and
       a2': "s= Normal s'" 
    using a0 a2 unfolding min_call_def by auto  
  then have "(n-1,Γ,(Q,t)#cfg1) ∈  cptn_mod_nest_call" 
    using a1 a2 a2' elim_cptn_mod_nest_call_n_dec[of n Γ f s' cfg1 Q]
    by (metis (no_types, lifting) SmallStepCon.redex.simps(7) call_f_step_not_s_eq_t_false 
         cptn_elim_cases(2) cptn_eq_cptn_mod_set cptn_mod_nest_cptn_mod option.sel)   
  thus ?thesis
  proof -
    obtain nn :: "(('b, 'a, 'c, 'd) LanguageCon.com × ('b, 'c) xstate) list ⇒ 
                  ('a ⇒ ('b, 'a, 'c, 'd) LanguageCon.com option) ⇒ nat ⇒ nat" where
      "∀x0 x1 x2. (∃v3<x2. (v3, x1, x0) ∈ cptn_mod_nest_call) = 
                  (nn x0 x1 x2 < x2 ∧ (nn x0 x1 x2, x1, x0) ∈ cptn_mod_nest_call)"
      by moura
    then have f1: "∀n f ps. (¬ min_call n f ps ∨ (n, f, ps) ∈ cptn_mod_nest_call ∧ 
                            (∀na. ¬ na < n ∨ (na, f, ps) ∉ cptn_mod_nest_call)) ∧ 
                            (min_call n f ps ∨ (n, f, ps) ∉ cptn_mod_nest_call ∨ 
                    nn ps f n < n ∧ (nn ps f n, f, ps) ∈ cptn_mod_nest_call)"
      by (meson min_call_def)
    then have f2: "(n, Γ, (P, s) # (Q, t) # cfg1) ∈ cptn_mod_nest_call ∧ 
                   (∀na. ¬ na < n ∨ (na, Γ, (P, s) # (Q, t) # cfg1) ∉ cptn_mod_nest_call)"
      using a1 assms(1) by blast
    obtain bb :: 'b where
      f3: "s = Normal bb"
      using a2 by blast
    then have f4: "(LanguageCon.com.Call f, Normal bb) = (P, s)"
      using a2 by blast
    have f5: "n - 1 < n"
      using f2 by (metis (no_types) Suc_diff_Suc a2 diff_Suc_eq_diff_pred elim_cptn_mod_nest_call_n_greater_zero lessI minus_nat.diff_0)
    have f6: "(LanguageCon.com.Call f, Normal bb) = (P, s)"
      using f3 a2 by blast
    have f7: "Normal bb = t"
      using f4 f2 by (metis (no_types) SmallStepCon.redex.simps(7) a2 
                            call_f_step_not_s_eq_t_false cptn_elim_cases(2) 
                            cptn_eq_cptn_mod_set cptn_mod_nest_cptn_mod)
    have "(nn ((Q, t) # cfg1) Γ (n - 1), Γ, (Q, Normal bb) # cfg1) ∈ cptn_mod_nest_call ⟶ 
              (Suc (nn ((Q, t) # cfg1) Γ (n - 1)), Γ, 
            (LanguageCon.com.Call f, Normal bb) # (Q, Normal bb) # cfg1) ∈ cptn_mod_nest_call"
      using a2 cptn_mod_nest_call.CptnModNestCall by fastforce
    then show ?thesis
      using f7 f6 f5 f2 f1 ‹(n - 1, Γ, (Q, t) # cfg1) ∈ cptn_mod_nest_call› less_trans_Suc by blast
  qed  
  
qed

lemma redex_not_call_seq_catch:
 assumes a0:"redex P = Call f ∧ P≠Call f"          
 shows "∃p1 p2. P = Seq p1 p2 ∨ P = Catch p1 p2"
using a0 unfolding min_call_def
proof(induct P)
qed(fastforce+)

lemma skip_all_skip:
  assumes a0:"(Γ,cfg)∈cptn" and
          a1:"cfg = (Skip,s)#cfg1"
  shows "∀i<length cfg. fst(cfg!i) = Skip"
using a0 a1
proof(induct cfg1 arbitrary:cfg s)
  case Nil thus ?case by auto
next
  case (Cons x xs)
  then obtain s' where x:"x = (Skip,s')"
    by (metis CptnMod_elim_cases(1) cptn_eq_cptn_mod_set  stepc_elim_cases(1))
  moreover have cptn:"(Γ,x#xs)∈cptn"
    using Cons.prems(1) Cons.prems(2) cptn_dest_pair by blast
  moreover have 
    xs:"x # xs = (LanguageCon.com.Skip, s') # xs" using x by auto
  ultimately show ?case using Cons(1)[OF cptn xs] Cons(3)
    using diff_Suc_1 fstI length_Cons less_Suc_eq_0_disj nth_Cons' by auto 
qed

lemma skip_all_skip_throw:
  assumes a0:"(Γ,cfg)∈cptn" and
          a1:"cfg = (Throw,s)#cfg1"
  shows "∀i<length cfg. fst(cfg!i) = Skip ∨ fst(cfg!i) = Throw"
using a0 a1
proof(induct cfg1 arbitrary:cfg s)
  case Nil thus ?case by auto
next
  case (Cons x xs)
  then obtain s' where x:"x = (Skip,s') ∨ x = (Throw, s')"
    by (metis CptnMod_elim_cases(10) cptn_eq_cptn_mod_set)    
  then have cptn:"(Γ,x#xs)∈cptn"
    using Cons.prems(1) Cons.prems(2) cptn_dest_pair by blast
  show ?case using x
  proof 
    assume "x=(Skip,s')" thus ?thesis using skip_all_skip Cons(3)
      using cptn fstI length_Cons less_Suc_eq_0_disj nth_Cons' nth_Cons_Suc skip_all_skip 
      by fastforce 
  next
    assume x:"x=(Throw,s')"
    moreover have cptn:"(Γ,x#xs)∈cptn"
      using Cons.prems(1) Cons.prems(2) cptn_dest_pair by blast
    moreover have 
      xs:"x # xs = (LanguageCon.com.Throw, s') # xs" using x by auto
    ultimately show ?case using Cons(1)[OF cptn xs] Cons(3)
    using diff_Suc_1 fstI length_Cons less_Suc_eq_0_disj nth_Cons' by auto 
  qed  
qed

  
lemma skip_min_nested_call_0:
  assumes a0:"min_call n Γ cfg" and
          a1:"cfg = (Skip,s)#cfg1"
  shows "n=0"
proof -
  have asm0:"(n, Γ, cfg) ∈ cptn_mod_nest_call" and 
       asm1:"(∀m<n. (m, Γ, cfg) ∉ cptn_mod_nest_call)"  
       using a0 unfolding min_call_def by auto  
  show ?thesis using a1 asm0 asm1
  proof (induct cfg1 arbitrary: cfg s n)
    case Nil thus ?case
      using cptn_mod_nest_call.CptnModNestOne neq0_conv by blast
  next
    case (Cons x xs)
      then obtain Q s' where cfg:"cfg = (LanguageCon.com.Skip, s) # (Q,s') # xs" by force
      then have min_call:"min_call n Γ cfg" using Cons unfolding min_call_def by auto
      then have "(∀f. SmallStepCon.redex Skip ≠ LanguageCon.com.Call f)" by auto
      then have "min_call n Γ ((Q, s')#xs)" 
        using elim_cptn_mod_min_nest_call[OF min_call cfg] cfg
        by simp
      thus ?case using Cons cfg unfolding min_call_def
      proof -
        assume a1: "(n, Γ, (Q, s') # xs) ∈ cptn_mod_nest_call ∧ (∀m<n. (m, Γ, (Q, s') # xs) ∉ cptn_mod_nest_call)"
        have "LanguageCon.com.Skip = Q"
          by (metis (no_types) ‹(n, Γ, cfg) ∈ cptn_mod_nest_call› cfg cptn_dest1_pair cptn_if_cptn_mod cptn_mod_nest_cptn_mod fst_conv last.simps last_length length_Cons lessI not_Cons_self2 skip_all_skip)
        then show ?thesis
          using a1 by (meson Cons.hyps)
      qed      
  qed
qed

lemma throw_min_nested_call_0:
  assumes a0:"min_call n Γ cfg" and
          a1:"cfg = (Throw,s)#cfg1"
  shows "n=0"
proof -
  have asm0:"(n, Γ, cfg) ∈ cptn_mod_nest_call" and 
       asm1:"(∀m<n. (m, Γ, cfg) ∉ cptn_mod_nest_call)"  
       using a0 unfolding min_call_def by auto  
  show ?thesis using a1 asm0 asm1
  proof (induct cfg1 arbitrary: cfg s n)
    case Nil thus ?case
      using cptn_mod_nest_call.CptnModNestOne neq0_conv by blast
  next
    case (Cons x xs)
      then obtain s' where x:"x = (Skip,s') ∨ x = (Throw, s')"
         using CptnMod_elim_cases(10) cptn_eq_cptn_mod_set
         by (metis cptn_mod_nest_cptn_mod)      
      then obtain Q where cfg:"cfg = (LanguageCon.com.Throw, s) # (Q,s') # xs"
        using Cons  by force
      then have min_call:"min_call n Γ cfg" using Cons unfolding min_call_def by auto
      then have "(∀f. SmallStepCon.redex Skip ≠ LanguageCon.com.Call f)" by auto
      then have min_call':"min_call n Γ ((Q, s')#xs)" 
        using elim_cptn_mod_min_nest_call[OF min_call cfg] cfg
        by simp
      from x show ?case
      proof
        assume "x=(Skip,s')"
        thus ?thesis using skip_min_nested_call_0 min_call' Cons(2) cfg by fastforce
      next       
        assume "x=(Throw,s')"
        thus ?thesis using Cons(1,2) min_call' cfg unfolding min_call_def  
          by blast 
      qed      
  qed
qed



text {* function to calculate that there is not any subsequent where the nested call is n *}
definition cond_seq_1
where 
"cond_seq_1 n Γ c1 s xs c2 zs ys ≡ ((n,Γ, (c1, s)#xs) ∈ cptn_mod_nest_call ∧ 
                       fst(last((c1,s)#xs)) = Skip ∧
                        (n,Γ,((c2, snd(last ((c1, s)#xs)))#ys)) ∈ cptn_mod_nest_call ∧
                       zs=(map (lift c2) xs)@((c2, snd(last ((c1, s)#xs)))#ys))"

definition cond_seq_2
where
"cond_seq_2 n Γ c1 s xs c2 zs ys s' s'' ≡  s= Normal s'' ∧ 
                    (n,Γ, (c1, s)#xs) ∈ cptn_mod_nest_call ∧ 
                    fst(last ((c1, s)#xs)) = Throw ∧
                    snd(last ((c1, s)#xs)) = Normal s' ∧ 
                    (n,Γ,(Throw,Normal s')#ys) ∈ cptn_mod_nest_call ∧ 
                     zs=(map (lift c2) xs)@((Throw,Normal s')#ys)"

definition cond_catch_1
where 
"cond_catch_1 n Γ c1 s xs c2 zs ys ≡ ((n,Γ, (c1, s)#xs) ∈ cptn_mod_nest_call ∧ 
                       fst(last((c1,s)#xs)) = Skip ∧
                        (n,Γ,((Skip, snd(last ((c1, s)#xs)))#ys)) ∈ cptn_mod_nest_call ∧
                       zs=(map (lift_catch c2) xs)@((Skip, snd(last ((c1, s)#xs)))#ys))"

definition cond_catch_2
where
"cond_catch_2 n Γ c1 s xs c2 zs ys s' s'' ≡ s= Normal s'' ∧ 
                    (n,Γ, (c1, s)#xs) ∈ cptn_mod_nest_call ∧ 
                    fst(last ((c1, s)#xs)) = Throw ∧
                    snd(last ((c1, s)#xs)) = Normal s' ∧ 
                    (n,Γ,(c2,Normal s')#ys) ∈ cptn_mod_nest_call ∧ 
                     zs=(map (lift_catch c2) xs)@((c2,Normal s')#ys)"

fun biggest_nest_call :: "('s,'p,'f,'e)com ⇒ 
                         ('s,'f) xstate ⇒ 
                         (('s,'p,'f,'e) config) list ⇒ 
                         ('s,'p,'f,'e) body ⇒ 
                         nat ⇒ bool"
where
 "biggest_nest_call (Seq c1 c2) s zs Γ n  = 
   (if (∃xs. ((min_call n Γ ((c1,s)#xs)) ∧ (zs=map (lift c2) xs))) then
       let xsa = (SOME xs. (min_call n Γ ((c1,s)#xs)) ∧ (zs=map (lift c2) xs)) in
       (biggest_nest_call c1 s xsa Γ n)
    else if (∃xs ys. cond_seq_1 n Γ c1 s xs c2 zs ys) then
         let xsa = (SOME xs. ∃ys. cond_seq_1 n Γ c1 s xs c2 zs ys);
             ysa = (SOME ys. cond_seq_1 n Γ c1 s xsa c2 zs ys) in
         if (min_call n Γ ((c2, snd(last ((c1, s)#xsa)))#ysa)) then True
         else (biggest_nest_call c1 s xsa Γ n)            
   else let xsa = (SOME xs. ∃ys s' s''. cond_seq_2 n Γ c1 s xs c2 zs ys s' s'') in
           (biggest_nest_call c1 s xsa Γ n))"      
|"biggest_nest_call (Catch c1 c2) s zs Γ n  = 
  (if (∃xs. ((min_call n Γ ((c1,s)#xs)) ∧ (zs=map (lift_catch c2) xs))) then
    let xsa = (SOME xs. (min_call n Γ ((c1,s)#xs)) ∧ (zs=map (lift_catch c2) xs)) in
       (biggest_nest_call c1 s xsa Γ n)
    else if (∃xs ys. cond_catch_1 n Γ c1 s xs c2 zs ys) then
         let xsa = (SOME xs. ∃ys. cond_catch_1 n Γ c1 s xs c2 zs ys) in            
                 (biggest_nest_call c1 s xsa Γ n)
   else let xsa = (SOME xs. ∃ys s' s''. cond_catch_2 n Γ c1 s xs c2 zs ys s' s'');
            ysa = (SOME ys. ∃s' s''. cond_catch_2 n Γ c1 s xsa c2 zs ys s' s'') in
         if (min_call n Γ ((c2, snd(last ((c1, s)#xsa)))#ysa)) then True
         else (biggest_nest_call c1 s xsa Γ n))"
|"biggest_nest_call _ _ _ _ _ = False"


lemma min_call_less_eq_n:
  "(n,Γ, (c1, s)#xs) ∈ cptn_mod_nest_call ⟹     
   (n,Γ,(c2, snd(last ((c1, s)#xs)))#ys) ∈ cptn_mod_nest_call ⟹
   min_call p Γ ((c1, s)#xs) ∧ min_call q Γ ((c2, snd(last ((c1, s)#xs)))#ys) ⟹
   p≤n ∧ q≤n"
unfolding min_call_def
using le_less_linear by blast

lemma min_call_seq_less_eq_n':
  "(n,Γ, (c1, s)#xs) ∈ cptn_mod_nest_call ⟹     
   min_call p Γ ((c1, s)#xs)  ⟹
   p≤n"
unfolding min_call_def
using le_less_linear by blast

lemma min_call_seq2:
  "min_call n Γ ((Seq c1 c2,s)#zs) ⟹
   (n,Γ, (c1, s)#xs) ∈ cptn_mod_nest_call ⟹ 
    fst(last ((c1, s)#xs)) = Skip ⟹
   (n,Γ,(c2, snd(last ((c1, s)#xs)))#ys) ∈ cptn_mod_nest_call ⟹
    zs=(map (lift c2) xs)@((c2, snd(last ((c1, s)#xs)))#ys) ⟹
   min_call n Γ ((c1, s)#xs) ∨ min_call n Γ ((c2, snd(last ((c1, s)#xs)))#ys)
   "
proof -
  assume a0:"min_call n Γ ((Seq c1 c2,s)#zs)" and
         a1:"(n,Γ, (c1, s)#xs) ∈ cptn_mod_nest_call" and
         a2:"fst(last ((c1, s)#xs)) = Skip" and
         a3:"(n,Γ,(c2, snd(last ((c1, s)#xs)))#ys) ∈ cptn_mod_nest_call" and
         a4:"zs=(map (lift c2) xs)@((c2, snd(last ((c1, s)#xs)))#ys)"
  then obtain p q where min_calls: 
    "min_call p Γ ((c1, s)#xs) ∧ min_call q Γ ((c2, snd(last ((c1, s)#xs)))#ys)"
    using a1 a3 minimum_nest_call by blast
  then have p_q:"p≤n ∧ q≤n" using a0 a1  a3 a4 min_call_less_eq_n by blast
  {
    assume ass0:"p<n ∧ q <n"     
    then have "(p,Γ, (c1, s)#xs) ∈ cptn_mod_nest_call" and
              "(q,Γ,(c2, snd(last ((c1, s)#xs)))#ys) ∈ cptn_mod_nest_call"
      using min_calls unfolding min_call_def by auto
    then have ?thesis
    proof (cases "p≤q")
      case True 
      then have q_cptn_c1:"(q, Γ, (c1, s) # xs) ∈ cptn_mod_nest_call" 
        using cptn_mod_nest_mono min_calls unfolding min_call_def  
        by blast
      have q_cptn_c2:"(q, Γ, (c2, snd (last ((c1, s) # xs))) # ys) ∈ cptn_mod_nest_call"
       using min_calls unfolding min_call_def by auto
      then have "(q,Γ,((Seq c1 c2,s)#zs)) ∈cptn_mod_nest_call"
        using True min_calls a2 a4  CptnModNestSeq2[OF q_cptn_c1 a2 q_cptn_c2 a4] 
        by auto
      thus ?thesis using ass0 a0 unfolding min_call_def by auto
    next
      case False
      then have q_cptn_c1:"(p, Γ, (c1, s) # xs) ∈ cptn_mod_nest_call" 
        using  min_calls unfolding min_call_def  
        by blast
      have q_cptn_c2:"(p, Γ, (c2, snd (last ((c1, s) # xs))) # ys) ∈ cptn_mod_nest_call"
       using min_calls False unfolding min_call_def
       by (metis (no_types, lifting) cptn_mod_nest_mono2 not_less)
      then have "(p,Γ,((Seq c1 c2,s)#zs)) ∈cptn_mod_nest_call"
        using False min_calls a2 a4  CptnModNestSeq2[OF q_cptn_c1 a2 q_cptn_c2 a4] 
        by auto
      thus ?thesis using ass0 a0 unfolding min_call_def by auto
    qed
  }note l=this
  {
    assume ass0:"p≥n ∨ q ≥n" 
    then have ?thesis using p_q min_calls by fastforce
  }
  thus ?thesis using l by fastforce
qed

lemma min_call_seq3:
  "min_call n Γ ((Seq c1 c2,s)#zs) ⟹
   s= Normal s'' ⟹
   (n,Γ, (c1, s)#xs) ∈ cptn_mod_nest_call ⟹ 
    fst(last ((c1, s)#xs)) = Throw ⟹
    snd(last ((c1, s)#xs)) = Normal s' ⟹
   (n,Γ,(Throw, snd(last ((c1, s)#xs)))#ys) ∈ cptn_mod_nest_call ⟹
    zs=(map (lift c2) xs)@((Throw, snd(last ((c1, s)#xs)))#ys) ⟹
   min_call n Γ ((c1, s)#xs)
   "
proof -
  assume a0:"min_call n Γ ((Seq c1 c2,s)#zs)" and
         a0':"s= Normal s''" and
         a1:"(n,Γ, (c1, s)#xs) ∈ cptn_mod_nest_call" and
         a2:"fst(last ((c1, s)#xs)) = Throw" and
         a2':"snd(last ((c1, s)#xs)) = Normal s'" and
         a3:"(n,Γ,(Throw, snd(last ((c1, s)#xs)))#ys) ∈ cptn_mod_nest_call" and
         a4:"zs=(map (lift c2) xs)@((Throw, snd(last ((c1, s)#xs)))#ys)"
  then obtain p where min_calls: 
    "min_call p Γ ((c1, s)#xs) ∧ min_call 0 Γ ((Throw, snd(last ((c1, s)#xs)))#ys)"
    using a1 a3 minimum_nest_call throw_min_nested_call_0  by metis
  then have p_q:"p≤n ∧ 0≤n" using a0 a1  a3 a4 min_call_less_eq_n by blast
  {
    assume ass0:"p<n ∧ 0 <n"     
    then have "(p,Γ, (c1, s)#xs) ∈ cptn_mod_nest_call" and
              "(0,Γ,(Throw, snd(last ((c1, s)#xs)))#ys) ∈ cptn_mod_nest_call"
      using min_calls unfolding min_call_def by auto
    then have ?thesis
    proof (cases "p≤0")
      case True 
      then have q_cptn_c1:"(0, Γ, (c1, Normal s'') # xs) ∈ cptn_mod_nest_call" 
        using cptn_mod_nest_mono min_calls a0' unfolding min_call_def  
        by blast
      have q_cptn_c2:"(0, Γ, (Throw, snd (last ((c1, s) # xs))) # ys) ∈ cptn_mod_nest_call"
       using min_calls unfolding min_call_def by auto
      then have "(0,Γ,((Seq c1 c2,s)#zs)) ∈cptn_mod_nest_call"
        using True min_calls a2 a4 a2' a0'  CptnModNestSeq3[OF q_cptn_c1 ] 
        by auto
      thus ?thesis using ass0 a0 unfolding min_call_def by auto
    next
      case False
      then have q_cptn_c1:"(p, Γ, (c1, Normal s'') # xs) ∈ cptn_mod_nest_call" 
        using  min_calls a0' unfolding min_call_def  
        by blast
      have q_cptn_c2:"(p, Γ, (Throw, snd (last ((c1, s) # xs))) # ys) ∈ cptn_mod_nest_call"
       using min_calls False unfolding min_call_def
       by (metis (no_types, lifting) cptn_mod_nest_mono2 not_less)
      then have "(p,Γ,((Seq c1 c2,s)#zs)) ∈cptn_mod_nest_call"
        using False min_calls a2 a4 a0' a2'  CptnModNestSeq3[OF q_cptn_c1]
        by auto       
      thus ?thesis using ass0 a0 unfolding min_call_def by auto
    qed
  }note l=this
  {
    assume ass0:"p≥n ∨ 0 ≥n" 
    then have ?thesis using p_q min_calls by fastforce
  }
  thus ?thesis using l by fastforce
qed

lemma min_call_catch2:
  "min_call n Γ ((Catch c1 c2,s)#zs) ⟹   
   (n,Γ, (c1, s)#xs) ∈ cptn_mod_nest_call ⟹ 
    fst(last ((c1, s)#xs)) = Skip ⟹    
   (n,Γ,(Skip, snd(last ((c1, s)#xs)))#ys) ∈ cptn_mod_nest_call ⟹
    zs=(map (lift_catch c2) xs)@((Skip, snd(last ((c1, s)#xs)))#ys) ⟹
   min_call n Γ ((c1, s)#xs)
   "
proof -
  assume a0:"min_call n Γ ((Catch c1 c2,s)#zs)" and        
         a1:"(n,Γ, (c1, s)#xs) ∈ cptn_mod_nest_call" and
         a2:"fst(last ((c1, s)#xs)) = Skip" and        
         a3:"(n,Γ,(Skip, snd(last ((c1, s)#xs)))#ys) ∈ cptn_mod_nest_call" and
         a4:"zs=(map (lift_catch c2) xs)@((Skip, snd(last ((c1, s)#xs)))#ys)"
  then obtain p where min_calls: 
    "min_call p Γ ((c1, s)#xs) ∧ min_call 0 Γ ((Skip, snd(last ((c1, s)#xs)))#ys)"
    using a1 a3 minimum_nest_call skip_min_nested_call_0  by metis
  then have p_q:"p≤n ∧ 0≤n" using a0 a1  a3 a4 min_call_less_eq_n by blast
  {
    assume ass0:"p<n ∧ 0 <n"     
    then have "(p,Γ, (c1, s)#xs) ∈ cptn_mod_nest_call" and
              "(0,Γ,(Skip, snd(last ((c1, s)#xs)))#ys) ∈ cptn_mod_nest_call"
      using min_calls unfolding min_call_def by auto
    then have ?thesis
    proof (cases "p≤0")
      case True 
      then have q_cptn_c1:"(0, Γ, (c1, s) # xs) ∈ cptn_mod_nest_call" 
        using cptn_mod_nest_mono min_calls  unfolding min_call_def  
        by blast
      have q_cptn_c2:"(0, Γ, (Skip, snd (last ((c1, s) # xs))) # ys) ∈ cptn_mod_nest_call"
       using min_calls unfolding min_call_def by auto
      then have "(0,Γ,((Catch c1 c2,s)#zs)) ∈cptn_mod_nest_call"
        using True min_calls a2 a4    CptnModNestCatch2[OF q_cptn_c1 ] 
        by auto
      thus ?thesis using ass0 a0 unfolding min_call_def by auto
    next
      case False
      then have q_cptn_c1:"(p, Γ, (c1, s) # xs) ∈ cptn_mod_nest_call" 
        using  min_calls  unfolding min_call_def  
        by blast
      have q_cptn_c2:"(p, Γ, (Skip, snd (last ((c1, s) # xs))) # ys) ∈ cptn_mod_nest_call"
       using min_calls False unfolding min_call_def
       by (metis (no_types, lifting) cptn_mod_nest_mono2 not_less)
      then have "(p,Γ,((Catch c1 c2,s)#zs)) ∈cptn_mod_nest_call"
        using False min_calls a2 a4   CptnModNestCatch2[OF q_cptn_c1]
        by auto       
      thus ?thesis using ass0 a0 unfolding min_call_def by auto
    qed
  }note l=this
  {
    assume ass0:"p≥n ∨ 0 ≥n" 
    then have ?thesis using p_q min_calls by fastforce
  }
  thus ?thesis using l by fastforce
qed

lemma min_call_catch_less_eq_n:
  "(n,Γ, (c1, Normal s)#xs) ∈ cptn_mod_nest_call ⟹        
   (n,Γ,(c2, snd(last ((c1, Normal s)#xs)))#ys) ∈ cptn_mod_nest_call ⟹    
   min_call p Γ ((c1, Normal s)#xs) ∧ min_call q Γ ((c2, snd(last ((c1, Normal s)#xs)))#ys) ⟹
   p≤n ∧ q≤n"
unfolding min_call_def
using le_less_linear by blast

lemma min_call_catch3:
  "min_call n Γ ((Catch c1 c2,Normal s)#zs) ⟹
   (n,Γ, (c1, Normal s)#xs) ∈ cptn_mod_nest_call ⟹ 
    fst(last ((c1, Normal s)#xs)) = Throw ⟹
    snd(last ((c1, Normal s)#xs)) = Normal s' ⟹
   (n,Γ,(c2, snd(last ((c1, Normal s)#xs)))#ys) ∈ cptn_mod_nest_call ⟹
    zs=(map (lift_catch c2) xs)@((c2, snd(last ((c1, Normal s)#xs)))#ys) ⟹
   min_call n Γ ((c1, Normal s)#xs) ∨ min_call n Γ ((c2, snd(last ((c1, Normal s)#xs)))#ys)
   "
proof -
  assume a0:"min_call n Γ ((Catch c1 c2,Normal s)#zs)" and
         a1:"(n,Γ, (c1, Normal s)#xs) ∈ cptn_mod_nest_call" and
         a2:"fst(last ((c1, Normal s)#xs)) = Throw" and
         a2':"snd(last ((c1, Normal s)#xs)) = Normal s'" and
         a3:"(n,Γ,(c2, snd(last ((c1, Normal s)#xs)))#ys) ∈ cptn_mod_nest_call" and
         a4:"zs=(map (lift_catch c2) xs)@((c2, snd(last ((c1, Normal s)#xs)))#ys)"
  then obtain p q where min_calls: 
    "min_call p Γ ((c1, Normal s)#xs) ∧ min_call q Γ ((c2, snd(last ((c1, Normal s)#xs)))#ys)"
    using a1 a3 minimum_nest_call by blast
  then have p_q:"p≤n ∧ q≤n" 
    using a1 a2 a2' a3 a4 min_call_less_eq_n by blast
  {
    assume ass0:"p<n ∧ q <n"     
    then have "(p,Γ, (c1, Normal s)#xs) ∈ cptn_mod_nest_call" and
              "(q,Γ,(c2, snd(last ((c1, Normal s)#xs)))#ys) ∈ cptn_mod_nest_call"
      using min_calls unfolding min_call_def by auto
    then have ?thesis
    proof (cases "p≤q")
      case True 
      then have q_cptn_c1:"(q, Γ, (c1,Normal s) # xs) ∈ cptn_mod_nest_call" 
        using cptn_mod_nest_mono min_calls unfolding min_call_def  
        by blast
      have q_cptn_c2:"(q, Γ, (c2, snd (last ((c1, Normal s) # xs))) # ys) ∈ cptn_mod_nest_call"
       using min_calls unfolding min_call_def by auto
      then have "(q,Γ,((Catch c1 c2, Normal s)#zs)) ∈cptn_mod_nest_call"
        using True min_calls a2 a2' a4  CptnModNestCatch3[OF q_cptn_c1 a2 a2' q_cptn_c2 a4] 
        by auto
      thus ?thesis using ass0 a0 unfolding min_call_def by auto
    next
      case False
      then have q_cptn_c1:"(p, Γ, (c1, Normal s) # xs) ∈ cptn_mod_nest_call" 
        using  min_calls unfolding min_call_def  
        by blast
      have q_cptn_c2:"(p, Γ, (c2, snd (last ((c1, Normal s) # xs))) # ys) ∈ cptn_mod_nest_call"
       using min_calls False unfolding min_call_def
       by (metis (no_types, lifting) cptn_mod_nest_mono2 not_less)
      then have "(p,Γ,((Catch c1 c2,Normal s)#zs)) ∈cptn_mod_nest_call"
        using False min_calls a2 a4  CptnModNestCatch3[OF q_cptn_c1 a2 a2' q_cptn_c2 a4] 
        by auto
      thus ?thesis using ass0 a0 unfolding min_call_def by auto
    qed
  }note l=this
  {
    assume ass0:"p≥n ∨ q ≥n" 
    then have ?thesis using p_q min_calls by fastforce
  }
  thus ?thesis using l by fastforce
qed

lemma min_call_seq_c1_not_finish:
  "min_call n Γ cfg ⟹
   cfg = (LanguageCon.com.Seq P0 P1, s) # (Q, t) # cfg1 ⟹
   (n, Γ,(P0, s)#xs) ∈ cptn_mod_nest_call ⟹
   (Q, t) # cfg1 = map (lift P1) xs ⟹
   min_call  n Γ ((P0, s)#xs)
   "
proof -
  assume a0:"min_call n Γ cfg" and
        a1:" cfg = (LanguageCon.com.Seq P0 P1, s) # (Q, t) # cfg1" and
        a2:"(n, Γ,(P0, s)#xs) ∈ cptn_mod_nest_call" and
        a3:"(Q, t) # cfg1 = map (lift P1) xs" 
  then have "(n, Γ,(P0, s)#xs) ∈ cptn_mod_nest_call" using a2 by auto
  moreover have "∀m<n. (m, Γ,(P0, s)#xs) ∉ cptn_mod_nest_call"
  proof-
    {fix m
     assume ass:"m<n"
     {  assume ass1:"(m, Γ, (P0, s) # xs) ∈ cptn_mod_nest_call"
       then have "(m,Γ,cfg) ∈  cptn_mod_nest_call" 
         using a1 a3 CptnModNestSeq1[OF ass1] by auto
       then have False using ass a0 unfolding min_call_def by auto
     }
     then have "(m, Γ, (P0, s) # xs) ∉ cptn_mod_nest_call" by auto
    } then show ?thesis by auto
  qed
  ultimately show ?thesis unfolding min_call_def by auto
qed

lemma min_call_seq_not_finish:
  " min_call  n Γ ((P0, s)#xs) ⟹
   cfg = (LanguageCon.com.Seq P0 P1, s) #  cfg1 ⟹  
    cfg1 = map (lift P1) xs ⟹
   min_call n Γ cfg 
   "
proof -
  assume a0:"min_call  n Γ ((P0, s)#xs)" and
        a1:" cfg = (LanguageCon.com.Seq P0 P1, s) #  cfg1" and        
        a2:" cfg1 = map (lift P1) xs" 
  then have "(n, Γ,cfg) ∈ cptn_mod_nest_call" 
    using a0 a1 a2 CptnModNestSeq1[of n Γ P0 s xs "cfg1" P1] unfolding min_call_def 
    by auto
  moreover have "∀m<n. (m, Γ,cfg) ∉ cptn_mod_nest_call"
  proof-
    {fix m
     assume ass:"m<n"
     {  assume ass1:"(m, Γ, cfg) ∈ cptn_mod_nest_call"
       then have "(m,Γ,(P0, s)#xs) ∈  cptn_mod_nest_call" 
         using a1 a2 by (metis (no_types) Seq_P_Not_finish div_seq_nest) 
       then have False using ass a0 unfolding min_call_def by auto
     }
     then have "(m, Γ, cfg) ∉ cptn_mod_nest_call" by auto
    } then show ?thesis by auto
  qed
  ultimately show ?thesis unfolding min_call_def by auto
qed

lemma min_call_catch_c1_not_finish:
  "min_call n Γ cfg ⟹
   cfg = (LanguageCon.com.Catch P0 P1, s) # (Q, t) # cfg1 ⟹
   (n, Γ,(P0, s)#xs) ∈ cptn_mod_nest_call ⟹
   (Q, t) # cfg1 = map (lift_catch P1) xs ⟹
   min_call  n Γ ((P0, s)#xs)
   "
proof -
  assume a0:"min_call n Γ cfg" and
        a1:" cfg = (LanguageCon.com.Catch P0 P1, s) # (Q, t) # cfg1" and
        a2:"(n, Γ,(P0, s)#xs) ∈ cptn_mod_nest_call" and
        a3:"(Q, t) # cfg1 = map (lift_catch P1) xs" 
  then have "(n, Γ,(P0, s)#xs) ∈ cptn_mod_nest_call" using a2 by auto
  moreover have "∀m<n. (m, Γ,(P0, s)#xs) ∉ cptn_mod_nest_call"
  proof-
    {fix m
     assume ass:"m<n"
     {  assume ass1:"(m, Γ, (P0, s) # xs) ∈ cptn_mod_nest_call"
       then have "(m,Γ,cfg) ∈  cptn_mod_nest_call" 
         using a1 a3 CptnModNestCatch1[OF ass1] by auto
       then have False using ass a0 unfolding min_call_def by auto
     }
     then have "(m, Γ, (P0, s) # xs) ∉ cptn_mod_nest_call" by auto
    } then show ?thesis by auto
  qed
  ultimately show ?thesis unfolding min_call_def by auto
qed

lemma min_call_catch_not_finish:
  " min_call  n Γ ((P0, s)#xs) ⟹
   cfg = (LanguageCon.com.Catch P0 P1, s) #  cfg1 ⟹  
    cfg1 = map (lift_catch P1) xs ⟹
   min_call n Γ cfg 
   "
proof -
  assume a0:"min_call  n Γ ((P0, s)#xs)" and
        a1:" cfg = (Catch P0 P1, s) #  cfg1" and        
        a2:" cfg1 = map (lift_catch P1) xs" 
  then have "(n, Γ,cfg) ∈ cptn_mod_nest_call" 
    using a0 a1 a2 CptnModNestCatch1[of n Γ P0 s xs "cfg1" P1] unfolding min_call_def 
    by auto
  moreover have "∀m<n. (m, Γ,cfg) ∉ cptn_mod_nest_call"
  proof-
    {fix m
     assume ass:"m<n"
     {  assume ass1:"(m, Γ, cfg) ∈ cptn_mod_nest_call"
       then have "(m,Γ,(P0, s)#xs) ∈  cptn_mod_nest_call" 
         using a1 a2 by (metis (no_types) Catch_P_Not_finish div_catch_nest) 
       then have False using ass a0 unfolding min_call_def by auto
     }
     then have "(m, Γ, cfg) ∉ cptn_mod_nest_call" by auto
    } then show ?thesis by auto
  qed
  ultimately show ?thesis unfolding min_call_def by auto
qed

lemma seq_xs_no_empty: assumes
     seq:"seq_cond_nest ((Q,t)#cfg1) P1 xs P0 s s'' s' Γ n" and
     cfg:"cfg = (LanguageCon.com.Seq P0 P1, s) # (Q, t) # cfg1" and
     a0:"SmallStepCon.redex (LanguageCon.com.Seq P0 P1) = LanguageCon.com.Call f"
     shows"∃Q' xs'. Q=Seq Q' P1 ∧ xs=(Q',t)#xs'"
using seq
unfolding lift_def seq_cond_nest_def
proof
    assume "(Q, t) # cfg1 = map (λ(P, s). (LanguageCon.com.Seq P P1, s)) xs"
    thus ?thesis by auto
next
  assume "fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Skip ∧
        (∃ys. (n, Γ, (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
              (Q, t) # cfg1 =
              map (λ(P, s). (LanguageCon.com.Seq P P1, s)) xs @
              (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∨
        fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Throw ∧
        snd (last ((P0, s) # xs)) = Normal s' ∧
        s = Normal s'' ∧
        (∃ys. (n, Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn_mod_nest_call ∧
              (Q, t) # cfg1 =
              map (λ(P, s). (LanguageCon.com.Seq P P1, s)) xs @
              (LanguageCon.com.Throw, Normal s') # ys)"
  thus ?thesis
  proof 
    assume ass:"fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Skip ∧
        (∃ys. (n, Γ, (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
              (Q, t) # cfg1 =
              map (λ(P, s). (LanguageCon.com.Seq P P1, s)) xs @
              (P1, snd (((P0, s) # xs) ! length xs)) # ys)"
    show ?thesis 
    proof (cases xs)
      case Nil thus ?thesis using cfg a0 ass by auto
    next
      case (Cons xa xsa)
      then obtain a b where xa:"xa = (a,b)" by fastforce
      obtain pps :: "(('a, 'b, 'c, 'd) LanguageCon.com × ('a, 'c) xstate) list" where
            "(Q, t) # cfg1 = ((case (a, b) of (c, x) ⇒ (LanguageCon.com.Seq c P1, x)) # map (λ(c, y). 
                             (LanguageCon.com.Seq c P1, y)) xsa) @ 
                             (P1, snd (((P0, s) # xs) ! length xs)) # pps"
        using xa ass local.Cons by moura
       then show ?thesis
         by (simp add: xa local.Cons)
    qed      
  next
    assume ass:"fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Throw ∧
        snd (last ((P0, s) # xs)) = Normal s' ∧
        s = Normal s'' ∧
        (∃ys. (n, Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn_mod_nest_call ∧
              (Q, t) # cfg1 =
              map (λ(P, s). (LanguageCon.com.Seq P P1, s)) xs @
              (LanguageCon.com.Throw, Normal s') # ys)"
    thus ?thesis
    proof (cases xs)
      case Nil thus ?thesis using cfg a0 ass by auto
    next
      case (Cons xa xsa)
      then obtain a b where xa:"xa = (a,b)" by fastforce
      obtain pps :: "(('a, 'b, 'c, 'd) LanguageCon.com × ('a, 'c) xstate) list" where
        "(Q, t) # cfg1 = ((case (a, b) of (c, x) ⇒ (LanguageCon.com.Seq c P1, x)) # map (λ(c, y). 
              (LanguageCon.com.Seq c P1, y)) xsa) @ (LanguageCon.com.Throw, Normal s') # pps"
        using ass local.Cons xa by force
      then show ?thesis
        by (simp add: local.Cons xa)
    qed           
  qed
qed

lemma catch_xs_no_empty: assumes
     seq:"catch_cond_nest ((Q,t)#cfg1) P1 xs P0 s s'' s' Γ n" and
     cfg:"cfg = (LanguageCon.com.Catch P0 P1, s) # (Q, t) # cfg1" and
     a0:"SmallStepCon.redex (LanguageCon.com.Catch P0 P1) = LanguageCon.com.Call f"
     shows"∃Q' xs'. Q=Catch Q' P1 ∧ xs=(Q',t)#xs'"
using seq
unfolding lift_catch_def catch_cond_nest_def
proof
    assume "(Q, t) # cfg1 = map (λ(P, s). (LanguageCon.com.Catch P P1, s)) xs"
    thus ?thesis by auto
next
  assume "fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Throw ∧
    snd (last ((P0, s) # xs)) = Normal s' ∧
    s = Normal s'' ∧
    (∃ys. (n, Γ, (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
          (Q, t) # cfg1 = map (λ(P, s). (LanguageCon.com.Catch P P1, s)) xs @ 
                                          (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∨
    fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Skip ∧
    (∃ys. (n, Γ, (LanguageCon.com.Skip, snd (last ((P0, s) # xs))) # ys) ∈ cptn_mod_nest_call ∧
          (Q, t) # cfg1 =
          map (λ(P, s). (LanguageCon.com.Catch P P1, s)) xs @ 
                         (LanguageCon.com.Skip, snd (last ((P0, s) # xs))) # ys)"
  thus ?thesis
  proof 
    assume ass:"fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Throw ∧
                snd (last ((P0, s) # xs)) = Normal s' ∧
                s = Normal s'' ∧
                (∃ys. (n, Γ, (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
                  (Q, t) # cfg1 = map (λ(P, s). (LanguageCon.com.Catch P P1, s)) xs @ 
                                          (P1, snd (((P0, s) # xs) ! length xs)) # ys)"
    show ?thesis 
    proof (cases xs)
      case Nil thus ?thesis using cfg a0 ass by auto
    next
      case (Cons xa xsa)
      then obtain a b where xa:"xa = (a,b)" by fastforce
      obtain pps :: "(('a, 'b, 'c, 'd) LanguageCon.com × ('a, 'c) xstate) list" where
       "(Q, t) # cfg1 = ((case (a, b) of (c, x) ⇒ (LanguageCon.com.Catch c P1, x)) # 
            map (λ(c, y). (LanguageCon.com.Catch c P1, y)) xsa) @ 
                           (P1, snd (((P0, s) # xs) ! length xs)) # pps"
       using ass local.Cons xa by moura
     then show ?thesis
       by (simp add: local.Cons xa)
    qed     
  next
    assume ass:"fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Skip ∧
    (∃ys. (n, Γ, (LanguageCon.com.Skip, snd (last ((P0, s) # xs))) # ys) ∈ cptn_mod_nest_call ∧
          (Q, t) # cfg1 =
          map (λ(P, s). (LanguageCon.com.Catch P P1, s)) xs @ 
                         (LanguageCon.com.Skip, snd (last ((P0, s) # xs))) # ys)"
    thus ?thesis
    proof (cases xs)
      case Nil thus ?thesis using cfg a0 ass by auto
    next
      case (Cons xa xsa)
      then obtain a b where xa:"xa = (a,b)" by fastforce      
      obtain pps :: "(('a, 'b, 'c, 'd) LanguageCon.com × ('a, 'c) xstate) list" where
        "(Q, t) # cfg1 = ((case (a, b) of (c, x) ⇒ 
           (LanguageCon.com.Catch c P1, x)) # map (λ(c, y). 
             (LanguageCon.com.Catch c P1, y)) xsa) @ 
               (LanguageCon.com.Skip, snd (last ((P0, s) # xs))) # pps"
        using ass local.Cons xa by force
      then show ?thesis
        by (simp add: local.Cons xa)
    qed        
  qed
qed

lemma redex_call_cptn_mod_min_nest_call_gr_zero:
 assumes a0:"min_call n Γ cfg" and
         a1:"cfg = (P,s)#(Q,t)#cfg1" and
         a2:"redex P = Call f ∧  
             Γ f = Some bdy ∧ (∃sa. s=Normal sa) ∧ t=s" and
         a3:"Γ⊢c(P,s)→(Q,t)"
 shows "n>0"
using a0 a1 a2 a3
proof (induct P arbitrary: Q cfg1 cfg s t n)
  case (Call f1) thus ?case
   by (metis SmallStepCon.redex.simps(7) elim_cptn_mod_nest_call_n_greater_zero min_call_def option.distinct(1) stepc_Normal_elim_cases(9))
next
  case (Seq P0 P1) 
  then obtain xs s' s'' where 
          p0_cptn:"(n, Γ,(P0, s)#xs) ∈ cptn_mod_nest_call"  and 
          seq:"seq_cond_nest ((Q,t)#cfg1) P1 xs P0 s s'' s' Γ n"
  using div_seq_nest[of n Γ cfg] unfolding min_call_def by blast
  then obtain m where min:"min_call m Γ ((P0, s)#xs)"
    using minimum_nest_call by blast 
  have xs':"∃Q' xs'. Q=Seq Q' P1 ∧ xs=(Q',t)#xs'"
     using seq Seq seq_xs_no_empty by auto 
  then have "0<m" using Seq(1,5,6) min
    using SmallStepCon.redex.simps(4) stepc_elim_cases_Seq_Seq by fastforce
  thus ?case by (metis min min_call_def not_gr0 p0_cptn) 
next
  case (Catch P0 P1)
 then obtain xs s' s'' where 
          p0_cptn:"(n, Γ,(P0, s)#xs) ∈ cptn_mod_nest_call"  and 
          seq:"catch_cond_nest ((Q,t)#cfg1) P1 xs P0 s s'' s' Γ n"
  using div_catch_nest[of n Γ cfg] unfolding min_call_def by blast
  then obtain m where min:"min_call m Γ ((P0, s)#xs)"
    using minimum_nest_call by blast 
  obtain Q' xs' where xs':"Q=Catch Q' P1 ∧ xs=(Q',t)#xs'"
     using catch_xs_no_empty[OF seq Catch(4)] Catch by blast
  then have "0<m" using Catch(1,5,6) min
    using SmallStepCon.redex.simps(4) stepc_elim_cases_Catch_Catch by fastforce
  thus ?case by (metis min min_call_def not_gr0 p0_cptn)
qed(auto)

  

lemma elim_redex_call_cptn_mod_min_nest_call:
 assumes a0:"min_call n Γ cfg" and
         a1:"cfg = (P,s)#(Q,t)#cfg1" and
         a2:"redex P = Call f ∧  
             Γ f = Some bdy ∧ (∃sa. s=Normal sa) ∧ t=s " and
         a3:"biggest_nest_call P s ((Q,t)#cfg1) Γ n"  
 shows "min_call n Γ ((Q,t)#cfg1)"
using a0 a1 a2 a3 
proof (induct P arbitrary: Q cfg1 cfg s t n)  
  case Cond thus ?case by fastforce
next
  case (Seq P0 P1) 
  then obtain xs s' s'' where 
          p0_cptn:"(n, Γ,(P0, s)#xs) ∈ cptn_mod_nest_call"  and 
          seq:"seq_cond_nest ((Q,t)#cfg1) P1 xs P0 s s'' s' Γ n"
  using div_seq_nest[of n Γ cfg] unfolding min_call_def by blast
  
  show ?case using seq unfolding seq_cond_nest_def 
  proof
    assume ass:"(Q, t) # cfg1 = map (lift P1) xs"   
    then obtain Q' xs' where xs':"Q=Seq Q' P1 ∧ xs=(Q',t)#xs'"
      unfolding lift_def by fastforce
    then have ctpn_P0:"(P0, s) # xs = (P0, s) # (Q', t) # xs'" by auto
    then have min_p0:"min_call n Γ ((P0, s)#xs)"
      using min_call_seq_c1_not_finish[OF Seq(3) Seq(4) p0_cptn] ass by auto
    then have ex_xs:"∃xs. min_call n Γ ((P0, s)#xs) ∧ (Q, t) # cfg1 = map (lift P1) xs" 
      using ass by auto
    then have min_xs:"min_call n Γ ((P0, s)#xs) ∧ (Q, t) # cfg1 = map (lift P1) xs" 
      using min_p0 ass by auto
    have "xs= (SOME xs. (min_call n Γ ((P0, s)#xs) ∧ (Q, t) # cfg1 = map (lift P1) xs))"
    proof -
      have "∀xsa. min_call n Γ ((P0, s)#xsa) ∧ (Q, t) # cfg1 = map (lift P1) xsa ⟶ xsa = xs"
        using xs' ass by (metis map_lift_eq_xs_xs')
      thus ?thesis using min_xs some_equality by (metis (mono_tags, lifting))
    qed
    then have big:"biggest_nest_call P0 s ((Q', t) # xs') Γ n" 
      using biggest_nest_call.simps(1)[of P0 P1 s "((Q, t) # cfg1)" Γ n] 
            Seq(6) xs' ex_xs by auto         
    have reP0:"redex P0 = (Call f) ∧ Γ f = Some bdy ∧ 
              (∃saa. s = Normal saa) ∧ t = s " using Seq(5) xs' by auto    
    have min_call:"min_call n Γ ((Q', t) # xs')" 
       using Seq(1)[OF min_p0 ctpn_P0 reP0] big xs' ass by auto
    thus ?thesis using min_call_seq_not_finish[OF min_call] ass xs' by blast
  next
    assume ass:"fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Skip ∧
                  (∃ys. (n, Γ, (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
                  (Q, t) # cfg1 = map (lift P1) xs @ (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∨
                fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Throw ∧
                  snd (last ((P0, s) # xs)) = Normal s' ∧
                  s = Normal s'' ∧
                  (∃ys. (n, Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn_mod_nest_call ∧
                     (Q, t) # cfg1 = map (lift P1) xs @ (LanguageCon.com.Throw, Normal s') # ys)"
    {assume ass:"fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Skip ∧
            (∃ys. (n, Γ, (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
            (Q, t) # cfg1 = map (lift P1) xs @ (P1, snd (((P0, s) # xs) ! length xs)) # ys)"     
     have ?thesis 
     proof (cases xs)
       case Nil thus ?thesis using Seq ass by fastforce
     next
       case (Cons xa xsa)
       then obtain ys where 
         seq2_ass:"fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Skip ∧ 
          (n, Γ, (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
          (Q, t) # cfg1 = map (lift P1) (xa#xsa) @ (P1, snd (((P0, s) # xs) ! length xs)) # ys"
          using ass by auto 
        then obtain mq mp1 where 
         min_call_q:"min_call mq Γ ((P0, s) # xs)" and
         min_call_p1:"min_call mp1 Γ ((P1, snd (((P0, s) # xs) ! length xs)) # ys)"         
       using seq2_ass minimum_nest_call p0_cptn by fastforce               
       then have mp: "mq≤n ∧ mp1 ≤n"
         using seq2_ass min_call_less_eq_n[of n Γ P0 s xs P1 ys  mq mp1] 
             Seq(3,4) p0_cptn by (simp add: last_length)
       have min_call:"min_call n Γ ((P0, s) # xs) ∨ 
             min_call n Γ ((P1, snd (((P0, s) # xs) ! length xs)) # ys)"
         using seq2_ass min_call_seq2[of n Γ P0 P1 s "(Q, t) # cfg1" xs ys] 
             Seq(3,4) p0_cptn by (simp add: last_length local.Cons)       
       from seq2_ass obtain Q' where Q':"Q=Seq Q' P1 ∧ xa=(Q',t)"          
       unfolding lift_def   
         by (metis (mono_tags, lifting) fst_conv length_greater_0_conv 
             list.simps(3) list.simps(9) nth_Cons_0 nth_append prod.case_eq_if prod.collapse snd_conv)  
       then have q'_n_cptn:"(n,Γ,(Q',t)#xsa)∈cptn_mod_nest_call" using p0_cptn Q' Cons
         using elim_cptn_mod_nest_call_n by blast 
       show ?thesis
       proof(cases "mp1=n")
         case True
         then have "min_call n Γ ((P1, snd (((P0, s) # xs) ! length xs)) # ys)"
           using min_call_p1 by auto
         then have min_P1:"min_call n Γ ((P1, snd ((xa # xsa) ! length xsa)) # ys)"
           using Cons seq2_ass by fastforce         
         then have p1_n_cptn:"(n, Γ,  (Q, t) # cfg1) ∈ cptn_mod_nest_call"
           using Seq.prems(1) Seq.prems(2) elim_cptn_mod_nest_call_n min_call_def by blast
         also then have "(∀m<n. (m, Γ, (Q, t) # cfg1) ∉ cptn_mod_nest_call)" 
         proof-
         { fix m
           assume ass:"m<n" 
           { assume Q_m:"(m, Γ, (Q, t) # cfg1) ∈ cptn_mod_nest_call"             
            then have False using min_P1 ass Q' Cons unfolding min_call_def
            proof -
              assume a1: "(n, Γ, (P1, snd ((xa # xsa) ! length xsa)) # ys) ∈ cptn_mod_nest_call ∧ (∀m<n. (m, Γ, (P1, snd ((xa # xsa) ! length xsa)) # ys) ∉ cptn_mod_nest_call)"
              have f2: "∀n f ps. (n, f, ps) ∉ cptn_mod_nest_call ∨ (∀x c ca psa. ps ≠ (LanguageCon.com.Seq (c::('b, 'a, 'c,'d) LanguageCon.com) ca, x) # psa ∨ (∃ps b ba. (n, f, (c, x) # ps) ∈ cptn_mod_nest_call ∧ seq_cond_nest psa ca ps c x ba b f n))"
                using div_seq_nest by blast
              have f3: "(P1, snd (last ((Q', t) # xsa))) # ys = (P1, snd (((P0, s) # xs) ! length xs)) # ys"
                by (simp add: Q' last_length local.Cons)
              have "fst (last ((Q', t) # xsa)) = LanguageCon.com.Skip"
                by (metis (no_types) Q' last_ConsR last_length list.distinct(1) local.Cons seq2_ass)
              then show ?thesis
                using f3 f2 a1 by (metis (no_types) Cons_lift_append Q' Seq_P_Ends_Normal Q_m ass seq2_ass)
            qed
           } 
         } then show ?thesis by auto
         qed           
         ultimately show ?thesis unfolding min_call_def by auto
       next
         case False
         then have "mp1<n" using mp by auto
         then have not_min_call_p1_n:"¬ min_call n Γ ((P1, snd (last ((P0, s) # xs))) # ys)"
           using min_call_p1 last_length unfolding min_call_def by metis
         then have min_call:"min_call n Γ ((P0, s) # xs)" 
           using min_call last_length unfolding min_call_def by metis
         then have "(P0, s) # xs = (P0, s) # xa#xsa"
           using Cons by auto
         then have big:"biggest_nest_call P0 s (((Q',t))#xsa) Γ n"
         proof-
           have "¬(∃xs. min_call n Γ ((P0, s)#xs) ∧ (Q, t) # cfg1 = map (lift P1) xs)"
             using min_call seq2_ass Cons
            proof -
              have "min_call n Γ ((LanguageCon.com.Seq P0 P1, s) # (Q, t) # cfg1)"
                using Seq.prems(1) Seq.prems(2) by blast
              then show ?thesis
                by (metis (no_types) Seq_P_Not_finish append_Nil2 list.simps(3) 
                          local.Cons min_call_def same_append_eq seq seq2_ass)
            qed
            moreover have "∃xs ys. cond_seq_1 n Γ P0 s xs P1 ((Q, t) # cfg1) ys"
              using seq2_ass p0_cptn unfolding cond_seq_1_def 
              by (metis last_length local.Cons) 
            moreover have "(SOME xs. ∃ys. cond_seq_1 n Γ P0 s xs P1 ((Q, t) # cfg1) ys) = xs"  
            proof -
              let ?P = "λxsa. ∃ys. (n, Γ, (P0, s) # xsa) ∈ cptn_mod_nest_call ∧
                   fst (last ((P0, s) # xsa)) = LanguageCon.com.Skip ∧
                   (n, Γ, (P1, snd (last ((P0, s) # xsa))) # ys) ∈ cptn_mod_nest_call ∧
                   (Q, t) # cfg1 = map (lift P1) xsa @ (P1, snd (last ((P0, s) # xsa))) # ys"             
              have "(⋀x. ∃ys. (n, Γ, (P0, s) # x) ∈ cptn_mod_nest_call ∧
               fst (last ((P0, s) # x)) = LanguageCon.com.Skip ∧
               (n, Γ, (P1, snd (last ((P0, s) # x))) # ys) ∈ cptn_mod_nest_call ∧
               (Q, t) # cfg1 = map (lift P1) x @ (P1, snd (last ((P0, s) # x))) # ys ⟹
                   x = xs)"              
              by (metis Seq_P_Ends_Normal cptn_mod_nest_call.CptnModNestSeq2 seq)
              moreover have "∃ys. (n, Γ, (P0, s) # xs) ∈ cptn_mod_nest_call ∧
                   fst (last ((P0, s) # xs)) = LanguageCon.com.Skip ∧
                   (n, Γ, (P1, snd (last ((P0, s) # xs))) # ys) ∈ cptn_mod_nest_call ∧
                   (Q, t) # cfg1 = map (lift P1) xs @ (P1, snd (last ((P0, s) # xs))) # ys"
                using ass  p0_cptn by (simp add: last_length)               
              ultimately show ?thesis using some_equality[of ?P xs]
                 unfolding cond_seq_1_def by blast 
            qed
            moreover have "(SOME ys. cond_seq_1 n Γ P0 s xs P1 ((Q, t) # cfg1) ys) = ys"
            proof -
               let ?P = "λys. (n, Γ, (P0, s) # xs) ∈ cptn_mod_nest_call ∧
                   fst (last ((P0, s) # xs)) = LanguageCon.com.Skip ∧
                   (n, Γ, (P1, snd (last ((P0, s) # xs))) # ys) ∈ cptn_mod_nest_call ∧
                   (Q, t) # cfg1 = map (lift P1) xs @ (P1, snd (last ((P0, s) # xs))) # ys"
                have "(n, Γ, (P0, s) # xs) ∈ cptn_mod_nest_call ∧
                   fst (last ((P0, s) # xs)) = LanguageCon.com.Skip ∧
                   (n, Γ, (P1, snd (last ((P0, s) # xs))) # ys) ∈ cptn_mod_nest_call ∧
                   (Q, t) # cfg1 = map (lift P1) xs @ (P1, snd (last ((P0, s) # xs))) # ys"
                 using p0_cptn seq2_ass Cons   by (simp add: last_length) 
                then show ?thesis using some_equality[of ?P ys]
                 unfolding cond_seq_1_def by fastforce      
            qed
            ultimately have "biggest_nest_call P0 s xs Γ n"
              using not_min_call_p1_n Seq(6) 
                    biggest_nest_call.simps(1)[of P0 P1 s "(Q, t) # cfg1" Γ n]
              by presburger
            then show ?thesis  using Cons Q' by auto
          qed
          have C:"(P0, s) # xs = (P0, s) # (Q', t) # xsa" using Cons Q' by auto
          have reP0:"redex P0 = (Call f) ∧ Γ f = Some bdy ∧ 
            (∃saa. s = Normal saa) ∧ t = s" using Seq(5) Q' by auto
          then have min_call:"min_call n Γ ((Q', t) # xsa)" using Seq(1)[OF min_call C reP0 big]
            by auto
          have p1_n_cptn:"(n, Γ,  (Q, t) # cfg1) ∈ cptn_mod_nest_call"
            using Seq.prems(1) Seq.prems(2) elim_cptn_mod_nest_call_n min_call_def by blast
          also then have "(∀m<n. (m, Γ, (Q, t) # cfg1) ∉ cptn_mod_nest_call)" 
          proof-
           { fix m
             assume ass:"m<n" 
             { assume Q_m:"(m, Γ, (Q, t) # cfg1) ∈ cptn_mod_nest_call" 
               then obtain xsa' s1 s1' where 
                  p0_cptn:"(m, Γ,(Q', t)#xsa') ∈ cptn_mod_nest_call"  and
                  seq:"seq_cond_nest cfg1 P1 xsa' Q' t s1 s1' Γ m"
               using div_seq_nest[of m Γ "(Q, t) # cfg1"] Q' by blast
               then have "xsa=xsa'" 
                 using seq2_ass 
                 Seq_P_Ends_Normal[of cfg1 P1 xsa Q' t ys m Γ xsa' s1 s1'] Cons
                 by (metis Cons_lift_append Q' Q_m last.simps last_length list.inject list.simps(3)) 
               then have False using min_call p0_cptn ass unfolding min_call_def by auto 
             } 
           } then show ?thesis by auto qed
             
         ultimately show ?thesis unfolding min_call_def by auto
       qed    
     qed
    }note l=this
    {assume ass:"fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Throw ∧
             snd (last ((P0, s) # xs)) = Normal s' ∧
            s = Normal s'' ∧ (∃ys. (n, Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn_mod_nest_call ∧
          (Q, t) # cfg1 = map (lift P1) xs @ (LanguageCon.com.Throw, Normal s') # ys)"          
     have ?thesis
     proof (cases "Γ⊢c(LanguageCon.com.Seq P0 P1, s) → (Q,t)")
       case True 
       thus  ?thesis
       proof (cases xs)
          case Nil thus ?thesis using Seq ass by fastforce
       next
         case (Cons xa xsa)
         then obtain ys where 
           seq2_ass:"fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Throw ∧
             snd (last ((P0, s) # xs)) = Normal s' ∧
            s = Normal s'' ∧  (n, Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn_mod_nest_call ∧
           (Q, t) # cfg1 = map (lift P1) xs @ (LanguageCon.com.Throw, Normal s') # ys"
            using ass by auto 
         then have t_eq:"t=Normal s''" using Seq by fastforce
         obtain mq mp1 where 
           min_call_q:"min_call mq Γ ((P0, s) # xs)" and
           min_call_p1:"min_call mp1 Γ ((Throw, snd (((P0, s) # xs) ! length xs)) # ys)"         
         using seq2_ass minimum_nest_call p0_cptn by (metis last_length)
         then have mp1_zero:"mp1=0" by (simp add: throw_min_nested_call_0)
         then have min_call: "min_call n Γ ((P0, s) # xs)"  
           using seq2_ass min_call_seq3[of n Γ P0 P1 s "(Q, t) # cfg1" s'' xs s' ys]
             Seq(3,4) p0_cptn by (metis last_length)      
         have n_z:"n>0" using redex_call_cptn_mod_min_nest_call_gr_zero[OF Seq(3) Seq(4) Seq(5) True]
           by auto          
         from seq2_ass obtain Q' where Q':"Q=Seq Q' P1 ∧ xa=(Q',t)"          
           unfolding lift_def using Cons
          proof -
            assume a1: "⋀Q'. Q = LanguageCon.com.Seq Q' P1 ∧ xa = (Q', t) ⟹ thesis"
            have "(LanguageCon.com.Seq (fst xa) P1, snd xa) = ((Q, t) # cfg1) ! 0"
             using seq2_ass unfolding lift_def
              by (simp add: Cons case_prod_unfold)
            then show ?thesis
              using a1 by fastforce
          qed  
         have big_call:"biggest_nest_call P0 s ((Q',t)#xsa) Γ n"
         proof-
           have "¬(∃xs. min_call n Γ ((P0, s)#xs) ∧ (Q, t) # cfg1 = map (lift P1) xs)"
             using min_call seq2_ass Cons Seq.prems(1) Seq.prems(2)
           by (metis Seq_P_Not_finish append_Nil2 list.simps(3) min_call_def same_append_eq seq)
           moreover have "¬(∃xs ys. cond_seq_1 n Γ P0 s xs P1 ((Q, t) # cfg1) ys)" 
             using min_call seq2_ass p0_cptn Cons Seq.prems(1) Seq.prems(2) 
             unfolding cond_seq_1_def
            by (metis com.distinct(17) com.distinct(71) last_length 
                      map_lift_some_eq seq_and_if_not_eq(4))
           moreover have "(SOME xs. ∃ys s' s''. cond_seq_2 n Γ P0 s xs P1 ((Q, t) # cfg1) ys s' s'') = xs"
           proof-
             let ?P="λxsa. ∃ys s' s''. s= Normal s'' ∧ 
                    (n,Γ, (P0, s)#xs) ∈ cptn_mod_nest_call ∧ 
                    fst(last ((P0, s)#xs)) = Throw ∧
                    snd(last ((P0, s)#xs)) = Normal s' ∧ 
                    (n,Γ,(Throw,Normal s')#ys) ∈ cptn_mod_nest_call ∧ 
                     ((Q, t) # cfg1)=(map (lift P1) xs)@((Throw,Normal s')#ys)"
             have "(⋀x. ∃ys s' s''. s= Normal s'' ∧ 
                    (n,Γ, (P0, s)#x) ∈ cptn_mod_nest_call ∧ 
                    fst(last ((P0, s)#x)) = Throw ∧
                    snd(last ((P0, s)#x)) = Normal s' ∧ 
                    (n,Γ,(Throw,Normal s')#ys) ∈ cptn_mod_nest_call ∧ 
                     ((Q, t) # cfg1)=(map (lift P1) x)@((Throw,Normal s')#ys) ⟹
                    x=xs)" using map_lift_some_eq seq2_ass by fastforce
             moreover have "∃ys s' s''. s= Normal s'' ∧ 
                    (n,Γ, (P0, s)#xs) ∈ cptn_mod_nest_call ∧ 
                    fst(last ((P0, s)#xs)) = Throw ∧
                    snd(last ((P0, s)#xs)) = Normal s' ∧ 
                    (n,Γ,(Throw,Normal s')#ys) ∈ cptn_mod_nest_call ∧ 
                     ((Q, t) # cfg1)=(map (lift P1) xs)@((Throw,Normal s')#ys)" 
                using ass p0_cptn by (simp add: last_length Cons)
            ultimately show ?thesis using some_equality[of ?P xs]
                 unfolding cond_seq_2_def by blast
         qed
           ultimately have "biggest_nest_call P0 s xs Γ n"
            using  Seq(6) 
                  biggest_nest_call.simps(1)[of P0 P1 s "(Q, t) # cfg1" Γ n]
            by presburger
           then show ?thesis  using Cons Q' by auto
         qed         
         have min_call:"min_call n Γ ((Q',t)#xsa)" 
           using Seq(1)[OF min_call _ _ big_call] Seq(5) Cons Q' by fastforce   
         then have p1_n_cptn:"(n, Γ,  (Q, t) # cfg1) ∈ cptn_mod_nest_call"
            using Seq.prems(1) Seq.prems(2) elim_cptn_mod_nest_call_n min_call_def by blast   
         also then have "(∀m<n. (m, Γ, (Q, t) # cfg1) ∉ cptn_mod_nest_call)" 
          proof-
           { fix m
             assume ass:"m<n" 
             { assume Q_m:"(m, Γ, (Q, t) # cfg1) ∈ cptn_mod_nest_call" 
               then obtain xsa' s1 s1' where 
                  p0_cptn:"(m, Γ,(Q', t)#xsa') ∈ cptn_mod_nest_call"  and
                  seq:"seq_cond_nest cfg1 P1 xsa' Q' (Normal s'') s1 s1' Γ m"
               using div_seq_nest[of m Γ "(Q, t) # cfg1"] Q' t_eq by blast
               then have "xsa=xsa'" 
                 using seq2_ass 
                 Seq_P_Ends_Abort[of cfg1 P1 xsa s' ys Q' s'' m Γ xsa' s1 s1' ] Cons Q' Q_m
                 by (simp add: Cons_lift_append last_length t_eq)                 
               then have False using min_call p0_cptn ass unfolding min_call_def by auto 
             } 
           } then show ?thesis by auto qed          
         ultimately show ?thesis unfolding min_call_def by auto
       qed
     next
       case False 
       then have env:"Γ⊢c(LanguageCon.com.Seq P0 P1, s) →e (Q,t)" using Seq
         by (meson elim_cptn_mod_nest_step_c min_call_def)
       moreover then have Q:"Q=Seq P0 P1" using env_c_c' by blast        
       ultimately show ?thesis using Seq
        proof -
          obtain nn :: "(('b, 'a, 'c,'d) LanguageCon.com × ('b, 'c) xstate) list ⇒ 
                         ('a ⇒ ('b, 'a, 'c,'d) LanguageCon.com option) ⇒ nat ⇒ nat" where
            f1: "∀x0 x1 x2. (∃v3<x2. (v3, x1, x0) ∈ cptn_mod_nest_call) = (nn x0 x1 x2 < x2 ∧ (nn x0 x1 x2, x1, x0) ∈ cptn_mod_nest_call)"
            by moura
          have f2: "(n, Γ, (LanguageCon.com.Seq P0 P1, s) # (Q, t) # cfg1) ∈ cptn_mod_nest_call ∧ (∀n. ¬ n < n ∨ (n, Γ, (LanguageCon.com.Seq P0 P1, s) # (Q, t) # cfg1) ∉ cptn_mod_nest_call)"
            using local.Seq(3) local.Seq(4) min_call_def by blast
          then have "¬ nn ((Q, t) # cfg1) Γ n < n ∨ (nn ((Q, t) # cfg1) Γ n, Γ, (Q, t) # cfg1) ∉ cptn_mod_nest_call"
            using False env env_c_c'  not_func_redex_cptn_mod_nest_n_env 
            by (metis Seq.prems(1) Seq.prems(2) min_call_def)
          then show ?thesis
            using f2 f1 by (meson elim_cptn_mod_nest_call_n min_call_def)
        qed
     qed          
    }
    thus ?thesis using l ass by fastforce
  qed
next
  case (Catch P0 P1) 
then obtain xs s' s'' where 
          p0_cptn:"(n, Γ,(P0, s)#xs) ∈ cptn_mod_nest_call"  and 
          catch:"catch_cond_nest ((Q,t)#cfg1) P1 xs P0 s s'' s' Γ n"
  using div_catch_nest[of n Γ cfg] unfolding min_call_def by blast
  
  show ?case using catch unfolding catch_cond_nest_def 
  proof
    assume ass:"(Q, t) # cfg1 = map (lift_catch P1) xs"   
    then obtain Q' xs' where xs':"Q=Catch Q' P1 ∧ xs=(Q',t)#xs'"
      unfolding lift_catch_def by fastforce
    then have ctpn_P0:"(P0, s) # xs = (P0, s) # (Q', t) # xs'" by auto
    then have min_p0:"min_call n Γ ((P0, s)#xs)"
      using min_call_catch_c1_not_finish[OF Catch(3) Catch(4) p0_cptn] ass by auto
    then have ex_xs:"∃xs. min_call n Γ ((P0, s)#xs) ∧ (Q, t) # cfg1 = map (lift_catch P1) xs" 
      using ass by auto
    then have min_xs:"min_call n Γ ((P0, s)#xs) ∧ (Q, t) # cfg1 = map (lift_catch P1) xs" 
      using min_p0 ass by auto
    have "xs= (SOME xs. (min_call n Γ ((P0, s)#xs) ∧ (Q, t) # cfg1 = map (lift_catch P1) xs))"
    proof -
      have "∀xsa. min_call n Γ ((P0, s)#xsa) ∧ (Q, t) # cfg1 = map (lift_catch P1) xsa ⟶ xsa = xs"
        using xs' ass by (metis map_lift_catch_eq_xs_xs')
      thus ?thesis using min_xs some_equality by (metis (mono_tags, lifting))
    qed
    then have big:"biggest_nest_call P0 s ((Q', t) # xs') Γ n" 
      using biggest_nest_call.simps(2)[of P0 P1 s "((Q, t) # cfg1)" Γ n] 
            Catch(6) xs' ex_xs by auto          
    have reP0:"redex P0 = (Call f) ∧ Γ f = Some bdy ∧ 
              (∃saa. s = Normal saa) ∧ t = s " using Catch(5) xs' by auto    
    have min_call:"min_call n Γ ((Q', t) # xs')" 
       using Catch(1)[OF min_p0 ctpn_P0 reP0] big xs' ass by auto
    thus ?thesis using min_call_catch_not_finish[OF min_call] ass xs' by blast
  next
    assume ass:"fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Throw ∧
                snd (last ((P0, s) # xs)) = Normal s' ∧
                s = Normal s'' ∧
               (∃ys. (n, Γ, (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
                  (Q, t) # cfg1 = map (lift_catch P1) xs @ (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∨
                   fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Skip ∧
                  (∃ys. (n, Γ, (LanguageCon.com.Skip, snd (last ((P0, s) # xs))) # ys) ∈ cptn_mod_nest_call ∧
                (Q, t) # cfg1 = map (lift_catch P1) xs @ (LanguageCon.com.Skip, snd (last ((P0, s) # xs))) # ys)" 
    {assume ass:"fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Throw ∧
                snd (last ((P0, s) # xs)) = Normal s' ∧
                s = Normal s'' ∧
               (∃ys. (n, Γ, (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
                  (Q, t) # cfg1 = map (lift_catch P1) xs @ (P1, snd (((P0, s) # xs) ! length xs)) # ys)"     
     have ?thesis 
     proof (cases xs)
       case Nil thus ?thesis using Catch ass by fastforce
     next
       case (Cons xa xsa)
       then obtain ys where 
         catch2_ass:"fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Throw ∧
                snd (last ((P0, s) # xs)) = Normal s' ∧
                s = Normal s'' ∧
                (n, Γ, (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧
                (Q, t) # cfg1 = map (lift_catch P1) xs @ (P1, snd (((P0, s) # xs) ! length xs)) # ys"
          using ass by auto 
        then obtain mq mp1 where 
         min_call_q:"min_call mq Γ ((P0, s) # xs)" and
         min_call_p1:"min_call mp1 Γ ((P1, snd (((P0, s) # xs) ! length xs)) # ys)"         
       using catch2_ass minimum_nest_call p0_cptn by fastforce               
       then have mp: "mq≤n ∧ mp1 ≤n"
         using catch2_ass min_call_less_eq_n 
             Catch(3,4) p0_cptn by (metis last_length) 
       have min_call:"min_call n Γ ((P0, s) # xs) ∨ 
             min_call n Γ ((P1, snd (((P0, s) # xs) ! length xs)) # ys)"
         using catch2_ass min_call_catch3[of n Γ P0 P1 s'' "(Q, t) # cfg1" xs s' ys] 
             Catch(3,4) p0_cptn by (metis last_length)       
       from catch2_ass obtain Q' where Q':"Q=Catch Q' P1 ∧ xa=(Q',t)"          
       unfolding lift_catch_def
        proof -
          assume a1: "⋀Q'. Q = LanguageCon.com.Catch Q' P1 ∧ xa = (Q', t) ⟹ thesis"
          assume "fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Throw ∧ snd (last ((P0, s) # xs)) = Normal s' ∧ s = Normal s'' ∧ (n, Γ, (P1, snd (((P0, s) # xs) ! length xs)) # ys) ∈ cptn_mod_nest_call ∧ (Q, t) # cfg1 = map (λ(P, s). (LanguageCon.com.Catch P P1, s)) xs @ (P1, snd (((P0, s) # xs) ! length xs)) # ys"
          then have "(LanguageCon.com.Catch (fst xa) P1, snd xa) = ((Q, t) # cfg1) ! 0"
            by (simp add: local.Cons prod.case_eq_if)
          then show ?thesis
            using a1 by force
        qed            
       then have q'_n_cptn:"(n,Γ,(Q',t)#xsa)∈cptn_mod_nest_call" using p0_cptn Q' Cons
         using elim_cptn_mod_nest_call_n by blast 
       show ?thesis
       proof(cases "mp1=n")
         case True
         then have "min_call n Γ ((P1, snd (((P0, s) # xs) ! length xs)) # ys)"
           using min_call_p1 by auto
         then have min_P1:"min_call n Γ ((P1, snd ((xa # xsa) ! length xsa)) # ys)"
           using Cons catch2_ass by fastforce         
         then have p1_n_cptn:"(n, Γ,  (Q, t) # cfg1) ∈ cptn_mod_nest_call"
           using Catch.prems(1) Catch.prems(2) elim_cptn_mod_nest_call_n min_call_def by blast
         also then have "(∀m<n. (m, Γ, (Q, t) # cfg1) ∉ cptn_mod_nest_call)" 
         proof-
         { fix m
           assume ass:"m<n" 
           { assume Q_m:"(m, Γ, (Q, t) # cfg1) ∈ cptn_mod_nest_call"   
             then have t_eq_s:"t=Normal s''" using Catch catch2_ass by fastforce                      
            then obtain xsa' s1 s1' where 
                  p0_cptn:"(m, Γ,(Q', t)#xsa') ∈ cptn_mod_nest_call"  and
                  catch_cond:"catch_cond_nest cfg1 P1 xsa' Q' (Normal s'') s1 s1' Γ m"
              using Q_m div_catch_nest[of m Γ "(Q, t) # cfg1"] Q' by blast
            have fst:"fst (last ((Q', Normal s'') # xsa)) = LanguageCon.com.Throw" 
              using catch2_ass Cons Q' by (simp add: last_length  t_eq_s)
            have cfg:"cfg1 = map (lift_catch P1) xsa @ (P1, snd (last ((Q', Normal s'') # xsa))) # ys"
              using catch2_ass Cons Q' by (simp add: last_length  t_eq_s)
            have snd:"snd (last ((Q', Normal s'') # xsa)) = Normal s'"
              using catch2_ass Cons Q' by (simp add: last_length  t_eq_s)
            then have "xsa=xsa' ∧ 
                   (m, Γ, (P1, snd (((Q', Normal s'') # xsa) ! length xsa)) # ys) ∈ cptn_mod_nest_call" 
              using catch2_ass Catch_P_Ends_Normal[OF cfg fst snd catch_cond] Cons
              by auto 
            then have False using min_P1 ass Q' t_eq_s unfolding min_call_def by auto              
           } 
         } then show ?thesis by auto
         qed           
         ultimately show ?thesis unfolding min_call_def by auto
       next
         case False
         then have "mp1<n" using mp by auto
         then have not_min_call_p1_n:"¬ min_call n Γ ((P1, snd (last ((P0, s) # xs))) # ys)"
           using min_call_p1 last_length unfolding min_call_def by metis
         then have min_call:"min_call n Γ ((P0, s) # xs)" 
           using min_call last_length unfolding min_call_def by metis
         then have "(P0, s) # xs = (P0, s) # xa#xsa"
           using Cons by auto
         then have big:"biggest_nest_call P0 s (((Q',t))#xsa) Γ n"
         proof-
           have "¬(∃xs. min_call n Γ ((P0, s)#xs) ∧ (Q, t) # cfg1 = map (lift_catch P1) xs)"
             using min_call catch2_ass Cons
            proof -
              have "min_call n Γ ((Catch P0 P1, s) # (Q, t) # cfg1)"
                using Catch.prems(1) Catch.prems(2) by blast
              then show ?thesis
                by (metis (no_types) Catch_P_Not_finish append_Nil2 list.simps(3) 
                     same_append_eq catch catch2_ass)
            qed
            moreover have "¬(∃xs ys. cond_catch_1 n Γ P0 s xs P1 ((Q, t) # cfg1) ys)"
              unfolding cond_catch_1_def using catch2_ass 
              by (metis Catch_P_Ends_Skip LanguageCon.com.distinct(17) catch last_length)
            moreover have "∃xs ys. cond_catch_2 n Γ P0 s xs P1 ((Q, t) # cfg1) ys s' s''"
              using catch2_ass p0_cptn unfolding cond_catch_2_def last_length
              by metis 
            moreover have "(SOME xs. ∃ys s' s''. cond_catch_2 n Γ P0 s xs P1 ((Q, t) # cfg1) ys s' s'') = xs"  
            proof -
              let ?P = "λxsa. s = Normal s'' ∧
                              (n, Γ, (P0, s) # xsa) ∈ cptn_mod_nest_call ∧
                              fst (last ((P0, s) # xsa)) = LanguageCon.com.Throw ∧
                              snd (last ((P0, s) # xsa)) = Normal s' ∧
                               (n, Γ, (P1, Normal s') # ys) ∈ cptn_mod_nest_call ∧ 
                               (Q, t) # cfg1 = map (lift_catch P1) xsa @ (P1, Normal s') # ys"             
              have "(⋀x. ∃ys s' s''. s = Normal s'' ∧
                              (n, Γ, (P0, s) # x) ∈ cptn_mod_nest_call ∧
                              fst (last ((P0, s) # x)) = LanguageCon.com.Throw ∧
                              snd (last ((P0, s) # x)) = Normal s' ∧
                               (n, Γ, (P1, Normal s') # ys) ∈ cptn_mod_nest_call ∧ 
                               (Q, t) # cfg1 = map (lift_catch P1) x @ (P1, Normal s') # ys ⟹
                   x = xs)"              
              by (metis Catch_P_Ends_Normal catch)
              moreover have "∃ys. s = Normal s'' ∧
                              (n, Γ, (P0, s) # xs) ∈ cptn_mod_nest_call ∧
                              fst (last ((P0, s) # xs)) = LanguageCon.com.Throw ∧
                              snd (last ((P0, s) # xs)) = Normal s' ∧
                               (n, Γ, (P1, Normal s') # ys) ∈ cptn_mod_nest_call ∧ 
                               (Q, t) # cfg1 = map (lift_catch P1) xs @ (P1, Normal s') # ys"
                using ass  p0_cptn   by (metis (full_types) last_length )             
              ultimately show ?thesis using some_equality[of ?P xs]
                 unfolding cond_catch_2_def by blast 
            qed
            moreover have "(SOME ys. ∃s' s''. cond_catch_2 n Γ P0 s xs P1 ((Q, t) # cfg1) ys s' s'') = ys"
            proof -
               let ?P = "λysa. s = Normal s'' ∧
                              (n, Γ, (P0, s) # xs) ∈ cptn_mod_nest_call ∧
                              fst (last ((P0, s) # xs)) = LanguageCon.com.Throw ∧
                              snd (last ((P0, s) # xs)) = Normal s' ∧
                               (n, Γ, (P1, Normal s') # ysa) ∈ cptn_mod_nest_call ∧ 
                               (Q, t) # cfg1 = map (lift_catch P1) xs @ (P1, Normal s') # ysa"
                have "(⋀x.  ∃s' s''. s = Normal s'' ∧
                          (n, Γ, (P0, s) # xs) ∈ cptn_mod_nest_call ∧
                          fst (last ((P0, s) # xs)) = LanguageCon.com.Throw ∧
                          snd (last ((P0, s) # xs)) = Normal s' ∧
                          (n, Γ, (P1, Normal s') # x) ∈ cptn_mod_nest_call ∧ (Q, t) # cfg1 = map (lift_catch P1) xs @ (P1, Normal s') # x ⟹
                          x = ys)" using catch2_ass by auto 
                moreover have "s = Normal s'' ∧
                      (n, Γ, (P0, s) # xs) ∈ cptn_mod_nest_call ∧
                       fst (last ((P0, s) # xs)) = LanguageCon.com.Throw ∧
                       snd (last ((P0, s) # xs)) = Normal s' ∧
                      (n, Γ, (P1, Normal s') # ys) ∈ cptn_mod_nest_call ∧ 
                       (Q, t) # cfg1 = map (lift_catch P1) xs @ (P1, Normal s') # ys"
                using ass  p0_cptn by (metis (full_types) catch2_ass last_length p0_cptn)             
                ultimately show ?thesis using some_equality[of ?P ys]
                 unfolding cond_catch_2_def by blast
            qed            
            ultimately have "biggest_nest_call P0 s xs Γ n"
              using not_min_call_p1_n Catch(6) 
                    biggest_nest_call.simps(2)[of P0 P1 s "(Q, t) # cfg1" Γ n]
              by presburger
            then show ?thesis  using Cons Q' by auto
          qed
          have C:"(P0, s) # xs = (P0, s) # (Q', t) # xsa" using Cons Q' by auto
          have reP0:"redex P0 = (Call f) ∧ Γ f = Some bdy ∧ 
            (∃saa. s = Normal saa) ∧ t = s " using Catch(5) Q' by auto
          then have min_call:"min_call n Γ ((Q', t) # xsa)" using Catch(1)[OF min_call C reP0 big]
            by auto
          have p1_n_cptn:"(n, Γ,  (Q, t) # cfg1) ∈ cptn_mod_nest_call"
            using Catch.prems(1) Catch.prems(2) elim_cptn_mod_nest_call_n min_call_def by blast
          also then have "(∀m<n. (m, Γ, (Q, t) # cfg1) ∉ cptn_mod_nest_call)" 
          proof-
           { fix m
             assume ass:"m<n" 
             { assume Q_m:"(m, Γ, (Q, t) # cfg1) ∈ cptn_mod_nest_call" 
               then have t_eq_s:"t=Normal s''" using Catch catch2_ass by fastforce
               then obtain xsa' s1 s1' where 
                  p0_cptn:"(m, Γ,(Q', t)#xsa') ∈ cptn_mod_nest_call"  and
                  catch_cond:"catch_cond_nest cfg1 P1 xsa' Q' (Normal s'') s1 s1' Γ m"
               using Q_m div_catch_nest[of m Γ "(Q, t) # cfg1"] Q' by blast
               have fst:"fst (last ((Q', Normal s'') # xsa)) = LanguageCon.com.Throw" 
                 using catch2_ass Cons Q' by (simp add: last_length  t_eq_s)
              have cfg:"cfg1 = map (lift_catch P1) xsa @ (P1, snd (last ((Q', Normal s'') # xsa))) # ys"
                 using catch2_ass Cons Q' by (simp add: last_length  t_eq_s)
              have snd:"snd (last ((Q', Normal s'') # xsa)) = Normal s'"
                using catch2_ass Cons Q' by (simp add: last_length  t_eq_s)
               then have "xsa=xsa'" 
                 using catch2_ass Catch_P_Ends_Normal[OF cfg fst snd catch_cond] Cons
                 by auto 
               then have False using min_call p0_cptn ass unfolding min_call_def by auto 
             } 
           } then show ?thesis by auto qed
         ultimately show ?thesis unfolding min_call_def by auto
       qed    
     qed
    }note l=this
    {assume ass:"fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Skip ∧
             (∃ys. (n, Γ, (LanguageCon.com.Skip, snd (last ((P0, s) # xs))) # ys) ∈ cptn_mod_nest_call ∧
             (Q, t) # cfg1 = map (lift_catch P1) xs @ (LanguageCon.com.Skip, snd (last ((P0, s) # xs))) # ys)"
     have ?thesis
     proof (cases "Γ⊢c(Catch P0 P1, s) → (Q,t)")
       case True 
       thus  ?thesis
       proof (cases xs)
          case Nil thus ?thesis using Catch ass by fastforce
       next
         case (Cons xa xsa)
         then obtain ys where 
           catch2_ass:"fst (((P0, s) # xs) ! length xs) = LanguageCon.com.Skip ∧
             (n, Γ, (LanguageCon.com.Skip, snd (last ((P0, s) # xs))) # ys) ∈ cptn_mod_nest_call ∧
             (Q, t) # cfg1 = map (lift_catch P1) xs @ (LanguageCon.com.Skip, snd (last ((P0, s) # xs))) # ys"
            using ass by auto 
         then have t_eq:"t=s" using Catch by fastforce
         obtain mq mp1 where 
           min_call_q:"min_call mq Γ ((P0, s) # xs)" and
           min_call_p1:"min_call mp1 Γ ((Skip, snd (((P0, s) # xs) ! length xs)) # ys)"         
         using catch2_ass minimum_nest_call p0_cptn by (metis last_length)
         then have mp1_zero:"mp1=0" by (simp add: skip_min_nested_call_0)
         then have min_call: "min_call n Γ ((P0, s) # xs)"  
           using catch2_ass min_call_catch2[of n Γ P0 P1 s "(Q, t) # cfg1" xs ys]
             Catch(3,4) p0_cptn by (metis last_length)      
         have n_z:"n>0" using redex_call_cptn_mod_min_nest_call_gr_zero[OF Catch(3) Catch(4) Catch(5) True]
           by auto          
         from catch2_ass obtain Q' where Q':"Q=Catch Q' P1 ∧ xa=(Q',t)"          
           unfolding lift_catch_def using Cons
          proof -
            assume a1: "⋀Q'. Q = Catch Q' P1 ∧ xa = (Q', t) ⟹ thesis"
            have "(Catch (fst xa) P1, snd xa) = ((Q, t) # cfg1) ! 0"
             using catch2_ass unfolding lift_catch_def
              by (simp add: Cons case_prod_unfold)
            then show ?thesis
              using a1 by fastforce
          qed  
         have big_call:"biggest_nest_call P0 s ((Q',t)#xsa) Γ n"
         proof-
           have "¬(∃xs. min_call n Γ ((P0, s)#xs) ∧ (Q, t) # cfg1 = map (lift_catch P1) xs)"
             using min_call catch2_ass Cons
           proof -
             have "min_call n Γ ((Catch P0 P1, s) # (Q, t) # cfg1)"
               using Catch.prems(1) Catch.prems(2) by blast
             then show ?thesis
               by (metis (no_types) Catch_P_Not_finish append_Nil2 list.simps(3) 
                     same_append_eq catch catch2_ass)
           qed
           moreover have "(∃xs ys. cond_catch_1 n Γ P0 s xs P1 ((Q, t) # cfg1) ys)"
             using catch2_ass p0_cptn unfolding cond_catch_1_def last_length
             by metis
           moreover have "(SOME xs. ∃ys. cond_catch_1 n Γ P0 s xs P1 ((Q, t) # cfg1) ys) = xs"
           proof -
             let ?P = "λxsa. ∃ys. (n, Γ,(P0, s)#xs) ∈ cptn_mod_nest_call ∧ 
                            fst (last ((P0, s) # xs)) = LanguageCon.com.Skip ∧
                             (n, Γ, (LanguageCon.com.Skip, 
                                snd (last ((P0, s) # xsa))) # ys) ∈ cptn_mod_nest_call ∧
                             (Q, t) # cfg1 = map (lift_catch P1) xsa @ 
                               (LanguageCon.com.Skip, snd (last ((P0, s) # xsa))) # ys"
             have "⋀xsa. ∃ys. (n, Γ,(P0, s)#xsa) ∈ cptn_mod_nest_call ∧ 
                             fst (last ((P0, s) # xs)) = LanguageCon.com.Skip ∧
                             (n, Γ, (LanguageCon.com.Skip, 
                                snd (last ((P0, s) # xsa))) # ys) ∈ cptn_mod_nest_call ∧
                             (Q, t) # cfg1 = map (lift_catch P1) xsa @ 
                               (LanguageCon.com.Skip, snd (last ((P0, s) # xsa))) # ys ⟹
                           xsa = xs"
             using Catch_P_Ends_Skip catch  catch2_ass map_lift_catch_some_eq by fastforce  
             moreover have "∃ys. (n, Γ,(P0, s)#xs) ∈ cptn_mod_nest_call ∧
                               fst (last ((P0, s) # xs)) = LanguageCon.com.Skip ∧
                             (n, Γ, (LanguageCon.com.Skip, 
                                snd (last ((P0, s) # xs))) # ys) ∈ cptn_mod_nest_call ∧
                             (Q, t) # cfg1 = map (lift_catch P1) xs @ 
                               (LanguageCon.com.Skip, snd (last ((P0, s) # xs))) # ys" 
               using ass p0_cptn by (simp add: last_length) 
             ultimately show ?thesis using some_equality[of ?P xs]
                 unfolding cond_catch_1_def by blast 
           qed           
           ultimately have "biggest_nest_call P0 s xs Γ n"
            using  Catch(6) 
                  biggest_nest_call.simps(2)[of P0 P1 s "(Q, t) # cfg1" Γ n]
            by presburger
           then show ?thesis  using Cons Q' by auto
         qed         
         have min_call:"min_call n Γ ((Q',t)#xsa)" 
           using Catch(1)[OF min_call _ _  big_call] Catch(5) Cons Q' by fastforce   
         then have p1_n_cptn:"(n, Γ,  (Q, t) # cfg1) ∈ cptn_mod_nest_call"
            using Catch.prems(1) Catch.prems(2) elim_cptn_mod_nest_call_n min_call_def by blast   
         also then have "(∀m<n. (m, Γ, (Q, t) # cfg1) ∉ cptn_mod_nest_call)" 
          proof-
           { fix m
             assume ass:"m<n" 
             { assume Q_m:"(m, Γ, (Q, t) # cfg1) ∈ cptn_mod_nest_call" 
               then obtain xsa' s1 s1' where 
                  p0_cptn:"(m, Γ,(Q', t)#xsa') ∈ cptn_mod_nest_call"  and
                  seq:"catch_cond_nest cfg1 P1 xsa' Q' t s1 s1' Γ m"
               using div_catch_nest[of m Γ "(Q, t) # cfg1"] Q' t_eq by blast
               then have "xsa=xsa'" 
                 using catch2_ass 
                 Catch_P_Ends_Skip[of cfg1 P1 xsa Q' t ys xsa' s1 s1']  
                 Cons Q' Q_m 
                 by (simp add:  last_length)                 
               then have False using min_call p0_cptn ass unfolding min_call_def by auto 
             } 
           } then show ?thesis by auto qed          
         ultimately show ?thesis unfolding min_call_def by auto
       qed
     next
       case False 
       then have env:"Γ⊢c(Catch P0 P1, s) →e (Q,t)" using Catch
         by (meson elim_cptn_mod_nest_step_c min_call_def)
       moreover then have Q:"Q=Catch P0 P1" using env_c_c' by blast        
       ultimately show ?thesis using Catch
        proof -
          obtain nn :: "(('b, 'a, 'c,'d) LanguageCon.com × ('b, 'c) xstate) list ⇒ ('a ⇒ ('b, 'a, 'c,'d) LanguageCon.com option) ⇒ nat ⇒ nat" where
            f1: "∀x0 x1 x2. (∃v3<x2. (v3, x1, x0) ∈ cptn_mod_nest_call) = (nn x0 x1 x2 < x2 ∧ (nn x0 x1 x2, x1, x0) ∈ cptn_mod_nest_call)"
            by moura
          have f2: "(n, Γ, (LanguageCon.com.Catch P0 P1, s) # (Q, t) # cfg1) ∈ cptn_mod_nest_call ∧ (∀n. ¬ n < n ∨ (n, Γ, (LanguageCon.com.Catch P0 P1, s) # (Q, t) # cfg1) ∉ cptn_mod_nest_call)"
            using local.Catch(3) local.Catch(4) min_call_def by blast
          then have "¬ nn ((Q, t) # cfg1) Γ n < n ∨ (nn ((Q, t) # cfg1) Γ n, Γ, (Q, t) # cfg1) ∉ cptn_mod_nest_call"
            using False env env_c_c'  not_func_redex_cptn_mod_nest_n_env 
            by (metis Catch.prems(1) Catch.prems(2) min_call_def)
          then show ?thesis
            using f2 f1 by (meson elim_cptn_mod_nest_call_n min_call_def)
        qed
     qed   
    }
    thus ?thesis using l ass by fastforce
  qed   
qed (fastforce)+


lemma cptn_mod_nest_n_1:
  assumes a0:"(n,Γ,cfs) ∈  cptn_mod_nest_call" and
          a1:"cfs=(p,s)#cfs'" and
          a2:"¬ (min_call n Γ cfs)"
  shows "(n-1,Γ,cfs) ∈  cptn_mod_nest_call"
using a0 a1 a2 
by (metis (no_types, lifting) Suc_diff_1 Suc_leI cptn_mod_nest_mono less_nat_zero_code min_call_def not_less)

lemma cptn_mod_nest_tl_n_1:
  assumes a0:"(n,Γ,cfs) ∈  cptn_mod_nest_call" and
          a1:"cfs=(p,s)#(q,t)#cfs'" and
          a2:"¬ (min_call n Γ cfs)"
  shows "(n-1,Γ,(q,t)#cfs') ∈  cptn_mod_nest_call"
  using a0 a1 a2
by (meson elim_cptn_mod_nest_call_n cptn_mod_nest_n_1) 

lemma cptn_mod_nest_tl_not_min:
  assumes a0:"(n,Γ,cfg) ∈  cptn_mod_nest_call" and
          a1:"cfg=(p,s)#cfg'" and
          a2:"¬ (min_call n Γ cfg)"
  shows "¬ (min_call n Γ cfg')"
proof (cases cfg')
  case Nil 
  have "(Γ, []) ∉ cptn"
    using cptn.simps by auto
  then show ?thesis unfolding min_call_def
    using cptn_eq_cptn_mod_set cptn_mod_nest_cptn_mod local.Nil by blast  
next
  case (Cons xa cfga)
  then obtain q t where "xa = (q,t)" by fastforce
  then have "(n-1,Γ,cfg') ∈  cptn_mod_nest_call"
    using a0 a1 a2 cptn_mod_nest_tl_n_1 Cons by fastforce
  also then have "(n,Γ,cfg') ∈  cptn_mod_nest_call"
    using cptn_mod_nest_mono Nat.diff_le_self by blast
  ultimately show ?thesis unfolding min_call_def
    using a0 a2 min_call_def by force 
qed
  
  

definition cpn :: "nat ⇒ ('s,'p,'f,'e) body ⇒ ('s,'p,'f,'e) com ⇒ 
                  ('s,'f) xstate ⇒ (('s,'p,'f,'e) confs) set" 
where
 "cpn n Γ P s ≡ {(Γ1,l). l!0=(P,s) ∧ (n,Γ,l) ∈ cptn_mod_nest_call ∧ Γ1=Γ}"


lemma cptn_mod_same_n:
  assumes a0:"(Γ,cfs)∈ cptn_mod" 
  shows "∃n. (n,Γ,cfs) ∈  cptn_mod_nest_call"
proof -
 show ?thesis using cptn_mod_nest_mono cptn_mod_cptn_mod_nest
 by (metis a0 cptn_mod_nest_mono2 leI)
qed

lemma cptn_mod_same_n1:
  assumes a0:"(Γ,cfs)∈ cptn_mod" and 
          a1:"(Γ,cfs1)∈ cptn_mod"
  shows "∃n. (n,Γ,cfs) ∈  cptn_mod_nest_call ∧ (n,Γ,cfs1) ∈  cptn_mod_nest_call"
proof -
 show ?thesis using cptn_mod_nest_mono cptn_mod_cptn_mod_nest
 by (metis a0 a1 cptn_mod_nest_mono2 leI)
qed


lemma dropcptn_is_cptn1 [rule_format,elim!]:
  "∀j<length c. (n,Γ,c) ∈ cptn_mod_nest_call ⟶ (n,Γ, drop j c) ∈ cptn_mod_nest_call"
proof -
  {fix j
   assume "j<length c ∧ (n,Γ,c) ∈ cptn_mod_nest_call"
   then have "(n,Γ, drop j c) ∈ cptn_mod_nest_call" 
   proof(induction j arbitrary: c)
     case 0 then show ?case by auto
   next
     case (Suc j) 
     then obtain a b c' where "c=a#b#c'"
       by (metis Cons_nth_drop_Suc Suc_lessE drop_0 less_trans_Suc zero_less_Suc)
     then also have "j<length (b#c')" using Suc by auto
     ultimately moreover have "(n, Γ, drop j (b # c')) ∈ cptn_mod_nest_call" using elim_cptn_mod_nest_call_n[of n Γ c] Suc
       by (metis surj_pair) 
     ultimately show ?case by auto  
   qed
 } thus ?thesis by auto 
qed

   

subsection ‹Compositionality of the Semantics›

subsubsection ‹Definition of the conjoin operator›

definition same_length :: "('s,'p,'f,'e) par_confs ⇒ (('s,'p,'f,'e) confs) list ⇒ bool" where
  "same_length c clist ≡ (∀i<length clist. length(snd (clist!i))=length (snd c))"

lemma same_length_non_pair:
  assumes a1:"same_length c clist " and
          a2:"clist'=map (λx. snd x) clist"
  shows "(∀i <length clist'. length( (clist'!i))=length (snd c))"
using a1 a2 by (auto simp add: same_length_def)


definition same_state :: "('s,'p,'f,'e) par_confs ⇒ (('s,'p,'f,'e) confs) list ⇒ bool" where
  "same_state c clist ≡ (∀i <length clist. ∀j<length (snd c). snd((snd c)!j) = snd((snd (clist!i))!j))"

lemma same_state_non_pair:
  assumes a1:"same_state c clist " and
          a2:"clist'=map (λx. snd x) clist"
  shows "(∀i <length clist'. ∀j<length (snd c). snd((snd c)!j) = snd( (clist'!i)!j))"
using a1 a2 by (auto simp add: same_state_def)

definition same_program :: "('s,'p,'f,'e) par_confs ⇒ (('s,'p,'f,'e) confs) list ⇒ bool" where
  "same_program c clist ≡ (∀j<length (snd c). fst((snd c)!j) = map (λx. fst(nth (snd x) j)) clist)"

lemma same_program_non_pair:
  assumes a1:"same_program c clist " and
          a2:"clist'=map (λx. snd x) clist"
  shows "(∀j<length (snd c). fst((snd c)!j) = map (λx. fst(nth x j)) clist')"
using a1 a2 by (auto simp add: same_program_def)

definition same_functions :: "('s,'p,'f,'e) par_confs ⇒ (('s,'p,'f,'e) confs) list ⇒ bool" where
 "same_functions c clist ≡ ∀i <length clist. fst (clist!i) = fst c"

definition compat_label :: "('s,'p,'f,'e) par_confs ⇒ (('s,'p,'f,'e) confs) list ⇒ bool" where
  "compat_label c clist ≡ 
     (∀j. Suc j<length (snd c) ⟶ 
         ( ((fst c)⊢p((snd c)!j)  → ((snd c)!(Suc j))) ∧ 
            (∃i<length clist. 
               ((fst (clist!i))⊢c ((snd (clist!i))!j)  → ((snd (clist!i))!(Suc j))) ∧ 
            (∀l<length clist. 
               l≠i ⟶  (fst (clist!l))⊢c (snd (clist!l))!j  →e ((snd (clist!l))!(Suc j))  ))) ∨ 
         ((fst c)⊢p((snd c)!j)  →e ((snd c)!(Suc j)) ∧ 
          (∀i<length clist.  (fst (clist!i))⊢c (snd (clist!i))!j  →e ((snd (clist!i))!(Suc j))   )))"

lemma compat_label_tran_0:
 assumes assm1:"compat_label c clist ∧ length (snd c) > Suc 0" 
 shows  "((fst c)⊢p((snd c)!0)  → ((snd c)!(Suc 0))) ∨ 
      ((fst c)⊢p((snd c)!0)  →e ((snd c)!(Suc 0)))"    
  using assm1 unfolding compat_label_def
 by blast
  
  
definition conjoin :: "(('s,'p,'f,'e) par_confs) ⇒ (('s,'p,'f,'e) confs) list ⇒ bool"  ("_ ∝ _" [65,65] 64) where
  "c ∝ clist ≡ (same_length c clist) ∧ (same_state c clist) ∧ (same_program c clist) ∧ 
                (compat_label c clist) ∧ (same_functions c clist)"


lemma conjoin_same_length: 
   "c ∝ clist ⟹ ∀i < length (snd c). length (fst ((snd c)!i)) =  length clist"
proof (auto)
  fix i
  assume a1:"c ∝ clist"
  assume a2:"i < length (snd c)"
  then have "(∀j<length (snd c). fst((snd c)!j) = map (λx. fst(nth (snd x) j)) clist)"
    using a1 unfolding conjoin_def same_program_def by auto
  thus "length (fst (snd c ! i)) = length clist" by (simp add: a2)
qed

lemma "c ∝ clist ⟹
       i< length (snd c) ∧ j < length (snd c) ⟹  
       length (fst ((snd c)!i)) = length (fst ((snd c)!j))"
using conjoin_same_length by fastforce

lemma conjoin_same_length_i_suci:"c ∝ clist ⟹
       Suc i< length (snd c) ⟹
       length (fst ((snd c)!i)) = length (fst ((snd c)!(Suc i)))"
using conjoin_same_length by fastforce

lemma conjoin_same_program_i:
  "c ∝ clist ⟹
   j < length (snd c) ⟹
   i < length clist ⟹
   fst ((snd (clist!i))!j) = (fst ((snd c)!j))!i"
proof -
  assume a0:"c ∝ clist" and
         a1:"j < length (snd c)" and
         a2:"i < length clist"
  have "length (fst ((snd c)!j)) = length clist"
    using conjoin_same_length a0 a1 by fastforce
  also have "fst (snd c ! j) = map (λx. fst (snd x ! j)) clist"
    using a0 a1 unfolding conjoin_def same_program_def by fastforce
  ultimately show ?thesis using a2 by fastforce
qed

lemma conjoin_same_program_i_j:
  "c ∝ clist ⟹
   Suc j < length (snd c) ⟹
   ∀l< length clist. fst ((snd (clist!l))!j) = fst ((snd (clist!l))!(Suc j)) ⟹
   fst ((snd c)!j) = (fst ((snd c)!(Suc j)))"
proof -
  assume a0:"c ∝ clist" and
         a1:"Suc j < length (snd c)" and
         a2:"∀l< length clist. fst ((snd (clist!l))!j) = fst ((snd (clist!l))!(Suc j))"
  have "length (fst ((snd c)!j)) = length clist"
    using conjoin_same_length a0 a1 by fastforce
  then have "map (λx. fst (snd x ! j)) clist = map (λx. fst (snd x ! (Suc j))) clist"
    using a2 by (metis (no_types, lifting) in_set_conv_nth map_eq_conv) 
  moreover have "fst (snd c ! j) = map (λx. fst (snd x ! j)) clist"
    using a0 a1 unfolding conjoin_def same_program_def by fastforce
  moreover have "fst (snd c ! Suc j) = map (λx. fst (snd x ! Suc j)) clist"
    using a0 a1 unfolding conjoin_def same_program_def by fastforce
  ultimately show ?thesis by fastforce
qed

lemma conjoin_last_same_state:
  assumes a0: "(Γ,l)∝ clist" and
   a1: "i < length clist" and
   a2: "(snd (clist!i))≠[]"
   shows "snd (last (snd (clist!i))) = snd (last l)"
proof -
  have "length l = length (snd (clist!i))" 
    using a0 a1 unfolding conjoin_def same_length_def by fastforce
  also then have length_l:"length l ≠0" using a2 by fastforce
  ultimately have "last (snd (clist!i)) = (snd (clist!i))!((length l)-1)" 
    using a1 a2 
    by (simp add: last_conv_nth)
  thus ?thesis using length_l a0 a1 unfolding conjoin_def same_state_def
    by (simp add:  a2 last_conv_nth )      
qed

lemma list_eq_if [rule_format]: 
  "∀ys. xs=ys ⟶ (length xs = length ys) ⟶ (∀i<length xs. xs!i=ys!i)"
  by (induct xs) auto



lemma list_eq: "(length xs = length ys ∧ (∀i<length xs. xs!i=ys!i)) = (xs=ys)"
apply(rule iffI)
 apply clarify
 apply(erule nth_equalityI)
 apply simp+
done

lemma nth_tl: "⟦ ys!0=a; ys≠[] ⟧ ⟹ ys=(a#(tl ys))"
  by (cases ys) simp_all

lemma nth_tl_if [rule_format]: "ys≠[] ⟶ ys!0=a ⟶ P ys ⟶ P (a#(tl ys))"
  by (induct ys) simp_all

lemma nth_tl_onlyif [rule_format]: "ys≠[] ⟶ ys!0=a ⟶ P (a#(tl ys)) ⟶ P ys"
  by (induct ys) simp_all

lemma nth_tl_eq [rule_format]: "ys≠[] ⟶ ys!0=a ⟶ P (a#(tl ys)) = P ys"
  by (induct ys) simp_all

lemma nth_tl_pair: "⟦p=(u,ys); ys!0=a; ys≠[] ⟧ ⟹ p=(u,(a#(tl ys)))"
by (simp add: SmallStepCon.nth_tl)

lemma nth_tl_eq_Pair [rule_format]: "p=(u,ys) ⟶ ys≠[] ⟶ ys!0=a ⟶ P ((u,a#(tl ys))) = P (u,ys)"
  by (induct ys) simp_all



lemma tl_in_cptn: "⟦ (g,a#xs) ∈cptn; xs≠[] ⟧ ⟹ (g,xs)∈cptn"
  by (force elim: cptn.cases)


lemma tl_zero[rule_format]: 
  " Suc j<length ys ⟶ P (ys!Suc j) ⟶ P (tl(ys)!j)"
  by (simp add: List.nth_tl)

lemma tl_zero1[rule_format]:
  "Suc j<length ys ⟶P (tl(ys)!j) ⟶P (ys!Suc j)"
 by (simp add: List.nth_tl)

lemma tl_zero_eq [rule_format]:
  "Suc j<length ys ⟶ (P (tl(ys)!j) = P (ys!Suc j))"
by (simp add: List.nth_tl)

lemma tl_zero_eq' :
   "∀j. Suc j<length ys ⟶ (P (tl(ys)!j) = P (ys!Suc j))"
using tl_zero_eq by blast

lemma tl_zero_pair:"i < length ys ⟹ length ys = length zs ⟹
       Suc j < length (snd (ys!i)) ⟹
       snd (zs!i) = tl (snd (ys!i)) ⟹        
       P ((snd (ys!i))!(Suc j)) =
       P ((snd (zs!i))!j)"  
  by (simp add: tl_zero_eq)


lemma tl_zero_pair':"∀i < length ys. length ys = length zs ⟶
       Suc j < length (snd (ys!i)) ⟶
       snd (zs!i) = tl (snd (ys!i)) ⟶        
       (P ((snd (ys!i))!(Suc j)) =
       P ((snd (zs!i))!j))"  
using tl_zero_pair by blast

lemma tl_zero_pair2:"i < length ys ⟹ length ys = length zs ⟹
       Suc (Suc j) < length (snd (ys!i)) ⟹
       snd (zs!i) = tl (snd (ys!i)) ⟹        
       P ((snd (ys!i))!(Suc (Suc j))) ((snd (ys!i))!(Suc j))  =
       P ((snd (zs!i))!(Suc j)) ((snd (zs!i))!j)"  
  by (simp add: tl_zero_eq)

lemma tl_zero_pair2':"∀i < length ys. length ys = length zs ⟶
       Suc (Suc j) < length (snd (ys!i)) ⟶
       snd (zs!i) = tl (snd (ys!i)) ⟶        
       P ((snd (ys!i))!(Suc (Suc j))) ((snd (ys!i))!(Suc j))  =
       P ((snd (zs!i))!(Suc j)) ((snd (zs!i))!j)"  
using tl_zero_pair2  by blast

lemma tl_zero_pair21:"∀i < length ys. length ys = length zs ⟶
       Suc (Suc j) < length (snd (ys!i)) ⟶
       snd (zs!i) = tl (snd (ys!i)) ⟶        
       P  ((snd (ys!i))!(Suc j))  ((snd (ys!i))!(Suc (Suc j)))=
       P ((snd (zs!i))!j) ((snd (zs!i))!(Suc j)) "
by (metis SmallStepCon.nth_tl list.size(3) not_less0 nth_Cons_Suc)  

lemma tl_pair:"Suc (Suc j) < length l ⟹     
       l1 = tl l ⟹
       P (l!(Suc (Suc j))) (l!(Suc j)) =
       P (l1!(Suc j)) (l1!j)"
by (simp add: tl_zero_eq)

lemma list_as_map: 
  assumes 
     a1:"length clist > 0" and 
     a2: "xs = (map (λx. fst (hd x)) clist)" and
     a3: "ys = (map (λx. tl x) clist)" and
     a4: "∀i< length clist. length (clist!i) > 0" and     
     a5: "∀i < length clist. ∀j< length clist. ∀k<length  (clist!i).
           snd ((clist!i)!k) = snd ((clist!j)!k)" and
     a6: "∀i < length clist. ∀j< length clist. 
            length (clist!i) = length (clist!j)" 
     shows "clist = map (λi. (fst i,snd ((clist!0)!0))#snd i) (zip xs ys)"
proof-
  let ?clist'= "map (λi. (fst i,snd ((clist!0)!0))#snd i) (zip xs ys)"
  have lens:"length clist = length ?clist'"  using a2 a3 by auto   
  have "(∀i<length clist. clist ! i = ?clist' ! i)" 
  proof -
    {
      fix i    
      assume a11:"i<length clist"
      have xs_clist:"xs!i = fst (hd (clist!i))" using a2 a11 by auto
      have ys_clist:"ys!i = tl (clist ! i)" using a3 a11 by auto
      have snd_zero:"snd (hd (clist!i)) = snd ((clist!0)!0)" using a5 a4 
        by (metis (no_types, lifting) a1 a11 hd_conv_nth less_numeral_extra(3) list.size(3))
      then have "(λi. (fst i,snd ((clist!0)!0))#snd i) ((zip xs ys)!i) = clist !i"               
        proof -
          have f1: "length xs = length clist"
            using a2 length_map by blast
          have "¬ (0::nat) < 0"
            by (meson less_not_refl)
          thus ?thesis
            using f1 by (metis (lifting) a11 a3 a4 
                         fst_conv length_map list.exhaust_sel 
                         list.size(3) nth_zip prod.collapse 
                         snd_conv snd_zero xs_clist ys_clist)
        qed      
      then have "clist ! i = ?clist' ! i" using lens a11 by force
    } 
    thus ?thesis by auto    
  qed
  thus ?thesis using lens list_eq by blast
qed

lemma list_as_map': 
  assumes 
     a1:"length clist > 0" and 
     a2: "xs = (map (λx. hd x) clist)" and
     a3: "ys = (map (λx. tl x) clist)" and
     a4: "∀i< length clist. length (clist!i) > 0"
     shows "clist = map (λi. (fst i)#snd i) (zip xs ys)"
proof-
  let ?clist'= "map (λi.(fst i)#snd i) (zip xs ys)"
  have lens:"length clist = length ?clist'" using a2 a3 by auto  
  have "(∀i<length clist. clist ! i = ?clist' ! i)" 
  proof -
    {
      fix i    
      assume a11:"i<length clist"
      have xs_clist:"xs!i = hd (clist!i)" using a2 a11 by auto
      have ys_clist:"ys!i = tl (clist ! i)" using a3 a11 by auto 
      then have "(λi. fst i#snd i) ((zip xs ys)!i) = clist !i"  
        using xs_clist ys_clist a11 a2 a3 a4 by fastforce  
      then have "clist ! i = ?clist' ! i" using lens a11 by force
    } 
    thus ?thesis by auto    
  qed
  thus ?thesis using lens list_eq by blast
qed


lemma conjoin_tl: 
  assumes 
    a1: "(Γ,x#xs) ∝ ys" and
    a2:"zs = map (λi. (fst i, tl (snd i))) ys"    
   shows "(Γ,xs) ∝ zs"
proof -
  have s_p:"same_program (Γ,x#xs) ys" using a1 unfolding conjoin_def by simp
  have s_l:"same_length (Γ,x#xs) ys" using a1 unfolding conjoin_def by simp
  have "∀i<length zs. snd (zs!i) = tl (snd (ys!i))"
    by (simp add: a2)    
  {
    have "same_length (Γ,xs) zs" using a1 a2 unfolding conjoin_def 
     by (simp add: same_length_def)
  } moreover note same_len = this
  {    
    {
       fix j
       assume a11:"j<length (snd (Γ, xs))"                                                               
       then have fst_suc:"fst (snd (Γ, xs) ! j) = fst(snd (Γ,x#xs)! Suc j)"
         by auto       
       then have "fst (snd (Γ, xs) ! j) = map (λx. fst (snd x ! j)) zs" 
       proof -
         have s_l_y_z:"length ys = length zs" using a2 by fastforce
         have Suc_j_l_ys:"∀i < length ys. Suc j < length (snd (ys!i))" 
           using a11 s_l unfolding same_length_def by fastforce
         have tail: "∀i < length ys. snd (zs!i) = tl (snd (ys!i))" using a2 
           by fastforce                  
         then have l_xs_zs_eq:"length (fst (snd (Γ, xs) ! j)) = length zs"
            using fst_suc s_l_y_z s_p a11 unfolding same_program_def by auto         
         then have "∀i<length ys. 
           fst (snd (Γ, x#xs) ! Suc j)!i = fst (snd (ys!i) ! (Suc j))"
             using s_p a11 unfolding same_program_def by fastforce
         then have "∀i<length zs. 
           fst (snd (Γ, x#xs) ! Suc j)!i = fst (snd (zs!i) ! (j))"
           using Suc_j_l_ys tail s_l_y_z tl_zero_pair by metis
        then have "∀i<length zs. 
           fst (snd (Γ, xs) ! j)!i = map (λx. fst (snd x !  j)) zs!i"
          using fst_suc by auto
        also have "length (fst (snd (Γ, xs) ! j)) = 
                   length (map (λx. fst (snd x !  j)) zs) " 
          using l_xs_zs_eq by auto
        ultimately show ?thesis using  l_xs_zs_eq list_eq by metis
       qed                 
    }
    then have "same_program  (Γ,xs) zs"
    unfolding conjoin_def  same_program_def same_length_def     
    by blast    
  }moreover note same_prog = this
  {
    have "same_state  (Γ,xs) zs" 
    using a1 a2 unfolding conjoin_def same_length_def same_state_def
    apply auto
    by (metis (no_types, hide_lams) List.nth_tl Suc_less_eq diff_Suc_1 length_tl nth_Cons_Suc)    
  }moreover note same_sta = this
  {
    have "same_functions  (Γ,xs) zs" 
     using a1 a2 unfolding conjoin_def
     apply auto
     apply (simp add: same_functions_def)          
     done
  }moreover note same_fun = this
  { {
      fix j
      assume a11:"Suc j<length (snd (Γ, xs))"
      have s_l_y_z:"length ys = length zs" using a2 by fastforce
      have Suc_j_l_ys:"∀i < length ys. Suc (Suc j) < length (snd (ys!i))" 
        using a11 s_l unfolding same_length_def by fastforce
      have tail: "∀i < length ys. snd (zs!i) = tl (snd (ys!i))" using a2 
        by fastforce    
      have same_env: "∀i < length ys. (fst (ys!i)) = Γ"
        using a1 unfolding conjoin_def same_functions_def by auto
      have fst: "∀x. fst(Γ, x) = Γ" by auto
      then have fun_ys_eq_fun_zs: "∀i < length ys. (fst (ys!i)) = (fst (zs!i))"
        using same_env s_l_y_z
        proof -
          have "∀n. ¬ n < length ys ∨ fst (zs ! n) = fst (ys ! n)"
            by (simp add: a2)
          thus ?thesis
            by presburger
        qed
      have suc_j:"Suc (Suc j) < length (snd (Γ, x#xs))" using a11 by auto      
     then have or_compat:"( (Γ ⊢p((snd  (Γ, x#xs))!(Suc j))  → ((snd  (Γ, x#xs))!(Suc (Suc j)))) ∧ 
            (∃i<length ys. 
               ((fst (ys!i))⊢c ((snd (ys!i))!(Suc j))  → ((snd (ys!i))!(Suc (Suc j)))) ∧ 
            (∀l<length ys. 
               l≠i ⟶ (fst (ys!l))⊢c (snd (ys!l))!(Suc j)  →e ((snd (ys!l))!(Suc (Suc j)))  ))) ∨ 
            (Γ⊢p((snd  (Γ, x#xs))!(Suc j))  →e ((snd  (Γ, x#xs))!(Suc (Suc j))) ∧ 
            (∀i<length ys. (fst (ys!i))⊢c (snd (ys!i))!(Suc j)  →e ((snd (ys!i))!(Suc (Suc j)))))"
        using suc_j a1 same_env unfolding conjoin_def compat_label_def fst by auto
       then have 
         "( (fst (Γ, xs) ⊢p((snd  (Γ, xs))!(j))  → ((snd  (Γ,xs))!((Suc j)))) ∧ 
              (∃i<length zs. 
                 ((fst (zs!i))⊢c ((snd (zs!i))!( j))  → ((snd (zs!i))!( (Suc j)))) ∧ 
              (∀l<length zs. 
                 l≠i ⟶ (fst (zs!l))⊢c (snd (zs!l))!( j)  →e (( snd (zs!l))!( (Suc j)))  )))∨
               ((fst (Γ, xs)⊢p((snd  (Γ, xs))!(j))  →e ((snd  (Γ, xs))!((Suc j))) ∧ 
           (∀i<length zs. (fst (zs!i))⊢c (snd (zs!i))!(j)  →e ((snd (zs!i))!((Suc j)))   )))"
       proof 
         assume a21:"( (Γ ⊢p((snd  (Γ, x#xs))!(Suc j))  → ((snd  (Γ, x#xs))!(Suc (Suc j)))) ∧ 
              (∃i<length ys. 
                 ((fst (ys!i))⊢c ((snd (ys!i))!(Suc j))  → ((snd (ys!i))!(Suc (Suc j)))) ∧ 
              (∀l<length ys. 
                 l≠i ⟶ (fst (ys!l))⊢c (snd (ys!l))!(Suc j)  →e ((snd (ys!l))!(Suc (Suc j)))  )))" 
          then obtain i where 
              f1:"( (Γ ⊢p((snd  (Γ, x#xs))!(Suc j))  → ((snd  (Γ, x#xs))!(Suc (Suc j)))) ∧ 
              (i<length ys ∧ 
                 ((fst (ys!i))⊢c ((snd (ys!i))!(Suc j))  → ((snd (ys!i))!(Suc (Suc j)))) ∧ 
              (∀l<length ys. 
                 l≠i ⟶ (fst (ys!l))⊢c (snd (ys!l))!(Suc j)  →e ((snd (ys!l))!(Suc (Suc j)))  )))"       
           by auto 
          then have "( (Γ ⊢p((snd  (Γ, x#xs))!(Suc j))  → ((snd  (Γ, x#xs))!(Suc (Suc j)))) ∧ 
              (∃i<length ys. 
                 ((fst (ys!i))⊢c ((snd (zs!i))!( j))  → ((snd (zs!i))!( (Suc j)))) ∧ 
              (∀l<length ys. 
                 l≠i ⟶ (fst (ys!l))⊢c (snd (zs!l))!( j)  →e (( snd (zs!l))!( (Suc j)))  )))"
            proof -
                have f1: "Γ⊢p snd (Γ, x # xs) ! Suc j → snd (Γ, x # xs) ! Suc (Suc j) ∧ i < length ys ∧ fst (ys ! i)⊢c snd (ys ! i) ! Suc j → snd (ys ! i) ! Suc (Suc j) ∧ (∀n. (¬ n < length ys ∨ n = i) ∨ fst (ys ! n)⊢c snd (ys ! n) ! Suc j →e snd (ys ! n) ! Suc (Suc j))"
                  using f1 by blast
                have f2: "j < length (snd (Γ, xs))"
                  by (meson Suc_lessD a11)
                have f3: "∀n. ¬ n < length zs ∨ length (snd (zs ! n)) = length (snd (Γ, xs))"
                  using same_len same_length_def by blast
                have "∀n. ¬ n < length ys ∨ snd (zs ! n) = tl (snd (ys ! n))"
                  using tail by blast
                thus ?thesis
                  using f3 f2 f1 by (metis (no_types) List.nth_tl a11 s_l_y_z)
              qed           
             then have"( (Γ ⊢p((snd  (Γ, xs))!(j))  → ((snd  (Γ,xs))!((Suc j)))) ∧ 
              (∃i<length zs. 
                 ((fst (zs!i))⊢c ((snd (zs!i))!( j))  → ((snd (zs!i))!( (Suc j)))) ∧ 
              (∀l<length zs. 
                 l≠i ⟶ (fst (zs!l))⊢c (snd (zs!l))!( j)  →e (( snd (zs!l))!( (Suc j)))  )))"
              using same_env s_l_y_z fun_ys_eq_fun_zs by force
             then have"( (fst (Γ, xs) ⊢p((snd  (Γ, xs))!(j))  → ((snd  (Γ,xs))!((Suc j)))) ∧ 
              (∃i<length zs. 
                 ((fst (zs!i))⊢c ((snd (zs!i))!( j))  → ((snd (zs!i))!( (Suc j)))) ∧ 
              (∀l<length zs. 
                 l≠i ⟶ (fst (zs!l))⊢c (snd (zs!l))!( j)  →e (( snd (zs!l))!( (Suc j)))  )))"
             by auto
             thus ?thesis
             by auto
      next    
        assume a22:
            "(Γ⊢p((snd  (Γ, x#xs))!(Suc j))  →e ((snd  (Γ, x#xs))!(Suc (Suc j))) ∧ 
            (∀i<length ys. (fst (ys!i))⊢c (snd (ys!i))!(Suc j)  →e ((snd (ys!i))!(Suc (Suc j)))   ))"
        then have 
          "(Γ⊢p((snd  (Γ, x#xs))!(Suc j))  →e ((snd  (Γ, x#xs))!(Suc (Suc j))) ∧ 
           (∀i<length ys. (fst (ys!i))⊢c (snd (zs!i))!(j)  →e ((snd (zs!i))!((Suc j)))   ))"
        using Suc_j_l_ys tail s_l_y_z tl_zero_pair21 by metis
        then have
          "(Γ⊢p((snd  (Γ, xs))!(j))  →e ((snd  (Γ, xs))!((Suc j))) ∧ 
           (∀i<length zs. (fst (zs!i))⊢c (snd (zs!i))!(j)  →e ((snd (zs!i))!((Suc j)))   ))"
          using same_env s_l_y_z fun_ys_eq_fun_zs by fastforce
        thus ?thesis by auto 
      qed
    }
    then have "compat_label  (Γ,xs) zs"
    using compat_label_def by blast 
  } note same_label = this
  ultimately show ?thesis using conjoin_def by auto
qed



lemma clist_tail: 
  assumes 
    a1:"length xs = length clist" and
    a2: "ys = (map (λi. (Γ,(fst i,s)#snd i)) (zip xs clist))"
 shows "∀i < length ys. tl (snd (ys!i)) = clist!i "
using a1 a2
proof -   
   show ?thesis using a2
   by (simp add: a1)           
qed   


lemma clist_map: 
   assumes 
    a1:"length xs = length clist" 
   shows "clist = map ((λp. tl (snd p)) ∘ (λi. (Γ, (fst i, s) # snd i))) (zip xs clist)"
proof -
  have f1: "map snd (zip xs clist) = clist"
    using a1 map_snd_zip by blast
  have "map snd (zip xs clist) = map ((λp. tl (snd p)) ∘ (λp. (Γ, (fst p, s) # snd p))) (zip xs clist)"
    by simp
  thus ?thesis
    using f1 by presburger
qed


lemma clist_map1: 
   assumes 
    a1:"length xs = length clist"     
   shows "clist = map (λp. tl (snd p)) (map (λi. (Γ,(fst i,s)#snd i)) (zip xs clist))"
proof -
   have "clist = map ((λp. tl (snd p)) ∘ (λi. (Γ, (fst i, s) # snd i))) (zip xs clist)" 
   using a1  clist_map by fastforce
   thus ?thesis by auto
qed

lemma clist_map2:
     "(clist = map (λp. tl (snd p)) (l::('a ×'b list) list) ) ⟹
       clist = map (λp. (snd p)) (map (λp. (fst p, tl (snd p))) (l::('a ×'b list) list)) "
by auto

lemma map_snd:
   assumes a1: "y = map  (λx. f x) l"
   shows   "y=(map snd (map (λx. (g x, f x)) l)) "
by (simp add: assms)
 
lemmas map_snd_sym = map_snd[THEN sym]

lemma map_snd':
   shows   " map  (λx. f x) l=(map snd (map (λx. (g x, f x)) l)) "
by simp

lemma clist_snd:
 assumes a1: "(Γ, a # ys) ∝ map (λx. (fst x, tl (snd x)))
                    (map (λi. (Γ, (fst i, s) # snd i)) (zip xs clist))" and
         a2: "length clist > 0 ∧ length clist = length xs"
 shows "clist = (map snd
          (map (λx. (Γ, (fst x, snd (clist ! 0 ! 0)) # snd x))
            (zip (map (λx. fst (hd x)) clist) (map tl clist))))"
proof -
     let ?concat_zip = "(λi. (Γ, (fst i, s) # snd i))"  
     let ?clist_ext = "map ?concat_zip (zip xs clist)"
     let ?exec_run = "(xs, s) # a # ys"
     let ?exec = "(Γ,?exec_run)"
     let ?exec_ext = "map (λx. (fst x, tl (snd x))) ?clist_ext"
     let ?zip = "(zip (map (λx. fst (hd x)) clist) 
                         (map (λx. tl x) clist))"
  have Γ_all: "∀i < length ?clist_ext. fst (?clist_ext !i) = Γ"
       by auto       
  have len:"length xs = length clist" using a2 by auto
  then have len_clist_exec:
   "length clist = length ?exec_ext" 
   by fastforce    
  then have len_clist_exec_map:
    "length ?exec_ext = 
              length (map (λx. (Γ, (fst x,snd ((clist!0)!0))#snd x)) 
                          ?zip)" 
   by fastforce               
  then have clist_snd:"clist = map (λx. snd x) ?exec_ext"
    using clist_map1 [of xs clist Γ s] clist_map2 len by blast   
  then have clist_len_eq_ays: 
      "∀i < length clist. length( (clist!i))=length (snd (Γ,a#ys))"      
    using len  same_length_non_pair a1 conjoin_def
    by blast
  then have clist_gz:"∀i < length clist. length (clist!i) > 0" 
    by fastforce
  have clist_len_eq: 
     "∀i < length clist. ∀j < length clist. 
        length (clist ! i) = length (clist ! j)" 
    using clist_len_eq_ays by auto          
  have clist_same_state: 
    "∀i < length clist. ∀j< length clist. ∀k<length  (clist!i).
       snd ((clist!i)!k) = snd ((clist!j)!k)"
  proof -
    have 
      "(∀i <length clist. ∀j<length (snd (Γ, a # ys)). snd((snd (Γ, a # ys))!j) = snd( (clist!i)!j))"
      using len clist_snd conjoin_def a1 conjoin_def same_state_non_pair 
    by blast
    thus ?thesis using clist_len_eq_ays by (metis (no_types))
  qed      
  then have clist_map:
    "clist = map (λi. (fst i,snd ((clist!0)!0))#snd i) ?zip"
    using list_as_map a2 clist_gz clist_len_eq by blast      
  moreover have "map (λi. (fst i,snd ((clist!0)!0))#snd i) ?zip =
             map snd (map (λx. (Γ, (fst x, snd (clist ! 0 ! 0)) # snd x))
       (zip (map (λx. fst (hd x)) clist) (map tl clist)))"
  using map_snd' by auto
  ultimately show ?thesis by auto   
qed

lemma list_as_zip:
 assumes a1: "(Γ, a # ys) ∝ map (λx. (fst x, tl (snd x)))
                    (map (λi. (Γ, (fst i, s) # snd i)) (zip xs clist))" and
         a2: "length clist > 0 ∧ length clist = length xs"
 shows "  map (λx. (fst x, tl (snd x)))
                    (map (λi. (Γ, (fst i, s) # snd i)) (zip xs clist)) =
          map (λx. (Γ, (fst x,snd ((clist!0)!0))#snd x)) 
                       (zip (map (λx. fst (hd x)) clist) 
                         (map (λx. tl x) clist))"
proof -
     let ?concat_zip = "(λi. (Γ, (fst i, s) # snd i))"  
     let ?clist_ext = "map ?concat_zip (zip xs clist)"
     let ?exec_run = "(xs, s) # a # ys"
     let ?exec = "(Γ,?exec_run)"
     let ?exec_ext = "map (λx. (fst x, tl (snd x))) ?clist_ext"
     let ?zip = "(zip (map (λx. fst (hd x)) clist) 
                         (map (λx. tl x) clist))"
  have Γ_all: "∀i < length ?clist_ext. fst (?clist_ext !i) = Γ"
       by auto       
  have len:"length xs = length clist" using a2 by auto
  then have len_clist_exec:
   "length clist = length ?exec_ext" 
   by fastforce    
  then have len_clist_exec_map:
    "length ?exec_ext = 
              length (map (λx. (Γ, (fst x,snd ((clist!0)!0))#snd x)) 
                          ?zip)" 
   by fastforce               
  then have clist_snd:"clist = map (λx. snd x) ?exec_ext"
    using clist_map1 [of xs clist Γ s] clist_map2 len by blast   
  then have clist_len_eq_ays: 
      "∀i < length clist. length( (clist!i))=length (snd (Γ,a#ys))"      
    using len  same_length_non_pair a1 conjoin_def
    by blast
  then have clist_gz:"∀i < length clist. length (clist!i) > 0" 
    by fastforce
  have clist_len_eq: 
     "∀i < length clist. ∀j < length clist. 
        length (clist ! i) = length (clist ! j)" 
    using clist_len_eq_ays by auto          
  have clist_same_state: 
    "∀i < length clist. ∀j< length clist. ∀k<length  (clist!i).
       snd ((clist!i)!k) = snd ((clist!j)!k)"
  proof -
    have 
      "(∀i <length clist. ∀j<length (snd (Γ, a # ys)). snd((snd (Γ, a # ys))!j) = snd( (clist!i)!j))"
      using len clist_snd conjoin_def a1 conjoin_def same_state_non_pair 
    by blast
    thus ?thesis using clist_len_eq_ays by (metis (no_types))
  qed      
  then have clist_map:
    "clist = map (λi. (fst i,snd ((clist!0)!0))#snd i) ?zip"
    using list_as_map a2 clist_gz clist_len_eq by blast      
  then have "∀i < length clist. 
                clist ! i = (fst (?zip!i),snd ((clist!0)!0)) # snd (?zip!i)"
  using len nth_map length_map by (metis (no_types, lifting))
  then have 
    "∀i < length clist. 
     ?exec_ext ! i = (Γ, (fst (?zip!i),snd ((clist!0)!0)) # snd (?zip!i))" 
  using Γ_all len  by fastforce           
  moreover have "∀i < length clist. 
    (Γ, (fst (?zip!i),snd ((clist!0)!0)) # snd (?zip!i)) = 
    (map (λx. (Γ, (fst x,snd ((clist!0)!0))#snd x)) 
                          ?zip)!i" 
  by auto        
  ultimately have 
     "∀i < length clist. 
       ?exec_ext ! i =(map (λx. (Γ, (fst x,snd ((clist!0)!0))#snd x)) 
                          ?zip)!i" 
  by auto
  then also have "length clist = length ?exec_ext" 
  using len by fastforce 
  ultimately have exec_ext_eq_clist_map:
     "∀i < length ?exec_ext. 
       ?exec_ext ! i =(map (λx. (Γ, (fst x,snd ((clist!0)!0))#snd x)) 
                          ?zip)!i" 
  by presburger
  then moreover have "length ?exec_ext = 
              length (map (λx. (Γ, (fst x,snd ((clist!0)!0))#snd x)) 
                          ?zip)" 
  using len clist_map by fastforce    
  ultimately show ?thesis
     using list_eq by blast  
qed

lemma hd_nth:
   assumes a1:"i< length l ∧ ( length( (l!i)) > 0)"
   shows "f (hd (l!i)) = f (nth (l!i) 0)"
using assms hd_conv_nth by fastforce

lemma map_hd_nth:
   assumes a1:"(∀i <length l. length( (l!i)) > 0)"
   shows "map (λx. f (hd x)) l = map (λx. f (nth (x) 0)) l"
proof -  
   have "∀i < length l. (map (λx. f (hd x)) l)!i = f (nth (l!i) 0)"
    using hd_nth a1 by auto
  moreover have "∀i < length l. (map (λx. f (nth x 0)) l)!i = f (nth (l!i) 0)"
    using hd_nth a1 by auto
  ultimately have f1:"∀i < length l. (map (λx. f (hd x)) l)!i =(map (λx. f (nth x 0)) l)!i "
    by auto
  moreover have f2:"length (map (λx. f (hd x)) l) = length l"
    by auto   
  moreover have "length (map (λx. f (nth x 0)) l) = length l" by auto
  ultimately show ?thesis using nth_equalityI by metis
qed

lemma "i<length clist ⟹ clist!i = (x1,ys) ⟹ ys = (map (λx. (fst (hd (snd x)),s)#tl (snd x)) clist)!i ⟹
         ys = (map (λx. (fst x, s)#snd x) 
               (zip (map (λx. fst (hd (snd x))) clist) 
                    (map (λx. tl (snd x)) clist)))!i"
proof (induct ys)
  case Nil thus ?case by auto
next
  case (Cons y ys) 
  have "∀n ps f. ¬ n < length ps ∨ map f ps ! n = (f (ps ! n::'a × ('b × 'c) list)::('b × 'c) list)"
    by force
  hence "y # ys = (fst (hd (snd (clist ! i))), s) # tl (snd (clist ! i))"
    using Cons.prems(1) Cons.prems(3) by presburger
  thus ?case
    using Cons.prems(1) by auto
qed

  

lemma clist_map_zip:"xs≠[] ⟹ (Γ,(xs,s)#ys) ∝ clist ⟹ 
      clist = map (λi. (Γ,(fst i,s)#(snd i))) (zip xs ((map (λx. tl (snd x))) clist))"
proof -
  let ?clist = "map snd clist"
  assume a1: "xs≠[]"
  assume a2:  "(Γ,(xs,s)#ys) ∝ clist"
  then have all_in_clist_not_empty:"∀i < length ?clist. (?clist!i) ≠ []"
   unfolding conjoin_def same_length_def by auto
  then have hd_clist:"∀i < length ?clist. hd (?clist!i) = (?clist!i)!0" 
     by (simp add: hd_conv_nth)  
  then have all_xs:"∀i< length ?clist. fst (hd (?clist!i)) = xs!i"
   using  a2 unfolding conjoin_def same_program_def by auto
  
  then have  all_s: "∀i < length ?clist. snd (hd (?clist!i)) = s"
    using a2 hd_clist unfolding conjoin_def same_state_def by fastforce
  have fst_clist_Γ:"∀i < length clist. fst (clist!i) = Γ"
    using a2 unfolding conjoin_def same_functions_def by auto 
   have p2:"length xs = length clist" using conjoin_same_length a2
   by fastforce
  
  
  then have "∀i< length (map (λx. fst (hd x)) ?clist). 
               (map (λx. fst (hd x)) ?clist)!i = xs!i"
    using all_xs by auto
  also have "length (map (λx. fst (hd x)) ?clist) = length xs" using p2 by auto
  ultimately have "(map (λx. fst (hd x)) ?clist) = xs"
   using nth_equalityI by metis
  then have xs_clist:"map (λx. fst (hd (snd x))) clist = xs" by auto
       
  have clist_hd_tl:"∀i < length ?clist. ?clist!i = hd (?clist!i) # (tl (?clist!i))"
   using all_in_clist_not_empty list.exhaust_sel by blast   
  then have "∀i < length ?clist. ?clist!i =(fst  (hd (?clist!i)),snd  (hd (?clist!i)))# (tl (?clist!i))"
    by auto
  then have "?clist = map (λx. (fst (hd x),snd (hd x))#tl x) ?clist" 
   using length_map list_eq_iff_nth_eq list_update_id map_update nth_list_update_eq
   by (metis (no_types, lifting) length_map list_eq_iff_nth_eq list_update_id map_update nth_list_update_eq)
  then have "?clist = map (λx. (fst (hd x),s)#tl x) ?clist"
   using all_s length_map nth_equalityI nth_map
    by (metis (no_types, lifting) ) 
  then have map_clist:"map (λx. (fst (hd (snd x)),s)#tl (snd x)) clist = ?clist" 
   by auto   
  then have "(map (λx. (fst x, s)#snd x) 
               (zip (map (λx. fst (hd (snd x))) clist) 
                    (map (λx. tl (snd x)) clist))) =  ?clist"     
    using map_clist  by (auto intro: nth_equalityI) 
  then have "∀i<length clist. clist!i =  (Γ,(map (λx. (fst x, s)#snd x) 
               (zip xs 
                   (map (λx. tl (snd x)) clist)))!i)" 
   using  xs_clist fst_clist_Γ  by auto   
  also have "length clist = length (map (λi. (Γ,(fst i,s)#(snd i))) (zip xs ((map (λx. tl (snd x))) clist)))" 
    using p2 by auto
  ultimately show "clist = map (λi. (Γ,(fst i,s)#(snd i))) (zip xs ((map (λx. tl (snd x))) clist))" 
    using length_map length_zip nth_equalityI nth_map 
    by (metis (no_types, lifting)) 
qed
            
lemma aux_if' : 
  assumes a:"length clist > 0 ∧ length clist = length xs ∧ 
             (∀i<length xs. (Γ,(xs!i,s)#clist!i) ∈ cptn) ∧ 
             ((Γ,(xs, s)#ys) ∝ map (λi. (Γ,(fst i,s)#snd i)) (zip xs clist))"
  shows "(Γ,(xs, s)#ys) ∈ par_cptn"
using a
proof (induct ys arbitrary: xs s clist) 
  case Nil then show ?case by (simp add: par_cptn.ParCptnOne)
next
  case (Cons a ys xs s clist)     
     let ?concat_zip = "(λi. (Γ, (fst i, s) # snd i))"  
     let ?com_clist_xs = "map ?concat_zip (zip xs clist)"
     let ?xs_a_ys_run = "(xs, s) # a # ys"
     let ?xs_a_ys_run_exec = "(Γ,?xs_a_ys_run)"
     let ?com_clist' = "map (λx. (fst x, tl (snd x))) ?com_clist_xs"
     let ?xs' = "(map (λx. fst (hd x)) clist)"     
     let ?clist' = "(map (λx. tl x) clist)"
     let ?zip_xs'_clist' = "zip ?xs' 
                            ?clist'"         
     obtain as sa where a_pair:"a=(as,sa)" by fastforce
       let ?comp_clist'_alt = "map (λx. (Γ, (fst x,snd ((clist!0)!0))#snd x)) ?zip_xs'_clist' "
       let ?clist'_alt = "map (λx. snd x) ?comp_clist'_alt"
       let ?comp_a_ys = "(Γ, (as,sa) # ys)"   
     have conjoin_hyp1:
       "(Γ, (as,sa) # ys) ∝ ?com_clist'"
       using conjoin_tl using a_pair Cons by blast     
     then have conjoin_hyp:
       "(Γ, (as,sa) # ys) ∝ map (λx. (Γ, (fst x,snd ((clist!0)!0))#snd x)) ?zip_xs'_clist'"
     using list_as_zip Cons.prems by fastforce    
     have len:"length xs = length clist" using Cons by auto 
     have clist_snd_map:
        "(map snd
          (map (λx. (Γ, (fst x, snd (clist ! 0 ! 0)) # snd x))
         (zip (map (λx. fst (hd x)) clist) (map tl clist)))) = clist"
       using clist_snd Cons.prems conjoin_hyp1 by fastforce
     have eq_len_clist_clist':
       "length ?clist' > 0" using Cons.prems by auto  
     have "(∀i <length clist. ∀j<length (snd ?comp_a_ys). snd((snd ?comp_a_ys)!j) = snd( (clist!i)!j))"
        using clist_snd_map conjoin_hyp conjoin_def same_state_non_pair[of ?comp_a_ys ?comp_clist'_alt ?clist'_alt]
         by fastforce   
     then have "∀i<length clist.
                  sa = snd ( (clist ! i)!0)" by fastforce
     also have clist_i_grz:"(∀i <length clist. length( (clist!i)) > 0)"
         using clist_snd_map conjoin_hyp conjoin_def same_length_non_pair[of ?comp_a_ys ?comp_clist'_alt ?clist'_alt]
     by fastforce  
     ultimately have all_i_sa_hd_clist:"∀i<length clist.
                  sa = snd (hd (clist ! i))"
     by (simp add: hd_conv_nth)      
     have as_sa_eq_xs'_s':"as = ?xs' ∧  sa = snd ((clist!0)!0)" 
     proof -              
       have "(∀j<length (snd ?comp_a_ys). fst((snd ?comp_a_ys)!j) = 
                map (λx. fst(nth x j)) ?clist'_alt)"       
       using conjoin_hyp conjoin_def same_program_non_pair[of ?comp_a_ys ?comp_clist'_alt ?clist'_alt]
       by fast
       then have are_eq:"fst((snd ?comp_a_ys)!0) = 
                map (λx. fst(nth x 0)) ?clist'_alt" by fastforce
       have fst_exec_is_as:"fst((snd ?comp_a_ys)!0) =as" by auto              
       then have "map (λx. fst(hd x)) clist=map (λx. fst(x!0)) clist"
         using map_hd_nth clist_i_grz by auto 
       then have "map (λx. fst(nth x 0)) ?clist'_alt =?xs'" using clist_snd_map map_hd_nth
        by fastforce
       moreover have "(∀i <length clist. ∀j<length (snd ?comp_a_ys). snd((snd ?comp_a_ys)!j) = snd( (clist!i)!j))"
        using clist_snd_map conjoin_hyp conjoin_def same_state_non_pair[of ?comp_a_ys ?comp_clist'_alt ?clist'_alt]
         by fastforce
       ultimately show ?thesis using are_eq fst_exec_is_as
          using Cons.prems by force 
    qed
    then have conjoin_hyp:
       "(Γ, (as,sa) # ys) ∝ map (λx. (Γ, (fst x,sa)#snd x))
                            (zip as (map tl clist))"
    using conjoin_hyp by auto
    then have eq_len_as_clist':
       "length as = length ?clist'" using Cons.prems as_sa_eq_xs'_s' by auto
    then have len_as_ys_eq:"length as = length xs" using Cons.prems by auto
    have " (∀i<length as. (Γ, ((as!i),sa)#(map (λx. tl x) clist)!i) ∈ cptn)" 
     using Cons.prems cptn_dest clist_snd_map len 
    proof -     
      have "∀i<length clist. clist!i = (hd (clist!i))#(tl (clist!i))" 
       using clist_i_grz 
      by auto
      then have "(∀i<length clist. (Γ, (xs ! i, s) # (hd (clist!i))#(tl (clist!i))) ∈ cptn)"
      using Cons.prems by auto
      then have f1:"(∀i<length clist. (Γ, (hd (clist!i))#(tl (clist!i))) ∈ cptn)"
      by (metis list.distinct(2) tl_in_cptn) 
      then have "(∀i<length clist. (Γ, ((as!i),sa)#(tl (clist!i))) ∈ cptn)"
      using as_sa_eq_xs'_s' all_i_sa_hd_clist by auto      
      then have "(∀i<length clist. (Γ, ((as!i),sa)#(map (λx. tl x) clist)!i) ∈ cptn)"
      by auto
      thus ?thesis using  len clist_i_grz len_as_ys_eq by auto
   qed
   then have a_ys_par_cptn:"(Γ, (as, sa) # ys) ∈ par_cptn"           
   using 
    conjoin_hyp eq_len_clist_clist' eq_len_as_clist'[THEN sym] Cons.hyps
   by blast  
   have Γ_all: "∀i < length ?com_clist_xs. fst (?com_clist_xs !i) = Γ"
   by auto
   have Gamma: "Γ= (fst ?xs_a_ys_run_exec)" by fastforce  
   have exec: "?xs_a_ys_run = (snd ?xs_a_ys_run_exec)" by fastforce  
   have split_par:
       "Γ⊢p ((xs, s) # a # ys) ! 0 → ((a # ys) ! 0) ∨
        Γ⊢p ((xs, s) # a # ys) ! 0 →e ((a # ys) ! 0)"     
       using compat_label_def compat_label_tran_0
             Cons.prems Gamma exec 
             compat_label_tran_0[of "(Γ, (xs, s) # a # ys)" 
                                   "(map (λi. (Γ, (fst i, s) # snd i)) (zip xs clist))"]    
       unfolding conjoin_def by auto      
     {
      assume "Γ⊢p ((xs, s) # a # ys) ! 0 → ((a # ys) ! 0)"      
      then have  " (Γ, (xs, s) # a # ys) ∈ par_cptn" 
      using a_ys_par_cptn a_pair par_cptn.ParCptnComp by fastforce
     } note env_sol=this
     {
      assume " Γ⊢p ((xs, s) # a # ys) ! 0 →e ((a # ys) ! 0)"      
      then have env_tran:" Γ⊢p (xs, s)  →e (as,sa)" using a_pair by auto
      have "xs = as"
       by (meson env_pe_c_c'_false env_tran)
      then have " (Γ, (xs, s) # a # ys) ∈ par_cptn" 
      using a_ys_par_cptn a_pair env_tran ParCptnEnv  by blast
     }
     then show "(Γ, (xs, s) # a # ys) ∈ par_cptn" using env_sol Cons split_par by fastforce
qed

lemma mapzip_upd:" length as = length clist  ⟹
       (map (λj. (as ! j, sa) # clist ! j) [0..<length as]) =  
       map (λj. ((fst j, sa)#snd j)) (zip as clist)"
proof -    
    assume a2: "length as = length clist"   
    have "∀i < length  (map (λj. (as ! j, sa) # clist ! j) [0..<length as]). (map (λj. (as ! j, sa) # clist ! j) [0..<length as])!i = map (λj. ((fst j, sa)#snd j)) (zip as clist)!i"  
     using a2
      by auto
  moreover have "length (map (λj. (as ! j, sa) # clist ! j) [0..<length as]) =
         length (map (λj. ((fst j, sa)#snd j)) (zip as clist))"
     using a2 by auto   
  ultimately have "(map (λj. (as ! j, sa) # clist ! j) [0..<length as]) = map (λj. ((fst j, sa)#snd j)) (zip as clist)"
     using nth_equalityI by blast
  thus "map (λj. (as ! j, sa) # clist ! j) [0..<length as] = 
        map (λj. (fst j, sa) # snd j) (zip as clist) "
      by auto
qed

lemma aux_if : 
  assumes a:" length clist = length xs ∧ 
             (∀i<length xs. (Γ,(xs!i,s)#clist!i) ∈ cptn) ∧ 
             ((Γ,(xs, s)#ys) ∝ map (λi. (Γ,(fst i,s)#snd i)) (zip xs clist))"
  shows "(Γ,(xs, s)#ys) ∈ par_cptn"
using a
proof (cases "length clist")
 case 0 
    then have clist_empty:"clist = []" by auto
    then have map_clist_empty:"map (λi. (Γ,(fst i,s)#snd i)) (zip xs clist) = []"
      by fastforce
    then have conjoin:"(Γ,(xs, s)#ys) ∝ []" using a by auto   
    then have all_eq:"∀j<length (snd (Γ,(xs, s)#ys)). fst (snd (Γ,(xs, s)#ys) ! j) = []"
      using conjoin_def same_program_def
    by (simp add: conjoin_def same_program_def)
    from conjoin          
    show ?thesis using conjoin
    proof (induct ys arbitrary: s xs) 
       case Nil then show ?case by (simp add: par_cptn.ParCptnOne)
    next
       case (Cons a ys)          
         then have conjoin_ind:"(Γ, (xs, s) # a # ys) ∝ []" by auto
         then have  "(Γ,(a # ys)) ∝ []" 
           by (auto simp add:conjoin_def same_length_def 
                 same_state_def same_program_def same_functions_def
                 compat_label_def)
         moreover obtain as sa where pair_a: "a=(as,sa)" using Cons by fastforce
         ultimately have ays_par_cptn:"(Γ, a # ys) ∈ par_cptn" using Cons.hyps by auto
         have "∀j. Suc j<length (snd (Γ,(xs, s)#(as,sa)#ys)) ⟶ 
                   ¬(∃i<length []. 
                     ((fst ([]!i))⊢c ((snd ([]!i))!j)  → ((snd ([]!i))!(Suc j))))"
         using conjoin_def compat_label_def by fastforce
         then have "(∀j. Suc j<length (snd (Γ,(xs, s)#(as,sa)#ys)) ⟶ 
                    ((fst (Γ,(xs, s)#(as,sa)#ys))⊢p((snd (Γ,(xs, s)#(as,sa)#ys))!j)  →e ((snd (Γ,(xs, s)#(as,sa)#ys))!(Suc j))))"
         using conjoin_def compat_label_def conjoin_ind pair_a by blast
         then have env_tran:"Γ⊢p (xs, s)  →e (as,sa)" by auto               
         then show " (Γ, (xs, s) # a # ys) ∈ par_cptn" 
         using ays_par_cptn pair_a env_tran ParCptnEnv env_pe_c_c'_false by blast
   qed
next
 case Suc
    then have "length clist > 0" by auto
    then show ?thesis using a aux_if' by blast
qed

lemma snormal_enviroment:"s = Normal nsa ∨ s = sa ∧ (∀sa. s ≠ Normal sa) ⟹
        Γ⊢c (x, s) →e (x, sa)"
by (metis Env Env_n)

lemma aux_onlyif [rule_format]: "∀xs s. (Γ,(xs, s)#ys) ∈ par_cptn ⟶ 
  (∃clist. (length clist = length xs) ∧ 
  (Γ, (xs, s)#ys) ∝ map (λi. (Γ, (fst i,s)#(snd i))) (zip xs clist) ∧ 
  (∀i<length xs. (Γ, (xs!i,s)#(clist!i)) ∈ cptn))"
proof (induct ys) 
  case Nil 
  {fix xs s
    assume "(Γ, [(xs, s)]) ∈ par_cptn"
    have f1:"length (map (λi. []) [0..<length xs]) = length xs" by auto
    have f2:"(Γ, [(xs, s)]) ∝ map (λi. (Γ, (fst i, s) # snd i))
                              (zip xs (map (λi. []) [0..<length xs]))"
    unfolding conjoin_def same_length_def same_functions_def same_state_def same_program_def compat_label_def            
      by(simp, rule nth_equalityI,simp,simp)
    note h = conjI[OF f1 f2]
    have f3:"(∀i<length xs. (Γ, (xs ! i, s) # (map (λi. []) [0..<length xs]) ! i) ∈ cptn)"
      by (simp add: cptn.CptnOne)
    note this = conjI[OF h f3]
    }
     thus ?case by blast
next
  case (Cons a ys) 
  {fix  xs s
   assume a1:"(Γ, (xs, s) # a # ys) ∈ par_cptn"
   then obtain as sa where a_pair: "a=(as,sa)" by fastforce
   then have par_cptn':"(Γ,( (as,sa)#ys)) ∈ par_cptn"
    using a1 par_cptn_dest by blast 
   then obtain clist where hyp: "
              length clist = length as ∧
              (Γ, (as, sa) #
                   ys) ∝ map (λi. (Γ, (fst i, sa) # snd i)) (zip as clist) ∧
              (∀i<length as. (Γ, (as ! i, sa) # clist ! i) ∈ cptn)"
     using Cons.hyps by fastforce
   have a11:"(Γ, (xs, s) # (as,sa) # ys) ∈ par_cptn" using a1 a_pair by auto
   have par_cptn_dest:"Γ⊢p (xs, s) →e (as, sa) ∨ Γ⊢p (xs, s) → (as, sa)"
     using par_cptn_elim_cases par_cptn' a1  a_pair by blast 
   {
     assume a1: "Γ⊢p (xs, s) →e (as, sa)"          
     then have xs_as_eq:"xs=as" by (meson env_pe_c_c'_false)
     then have ce:"∀i < length xs. Γ⊢c (xs!i, s) →e (as!i, sa)" using a1 pe_ce by fastforce
     let ?clist="(map (λj. (xs!j, sa)#(clist!j)) [0..<length xs])"    
     have s1:"length ?clist = length xs"
       by auto
     have s2:"(∀i<length xs. (Γ, (xs ! i, s) # ?clist ! i) ∈ cptn)"  
        using a1 hyp CptnEnv xs_as_eq ce by fastforce
     have s3:"(Γ, (xs, s) #
                       (as,sa) # ys) ∝ map (λi. (Γ, (fst i, s) # snd i))
                                   (zip xs ?clist)"     
     proof -        
         have s_len:"same_length (Γ, (xs, s) # (as,sa) # ys) 
                           (map (λi. (Γ, (fst i, s) # snd i))
                                   (zip xs ?clist))"
               using hyp conjoin_def same_length_def xs_as_eq a1 by fastforce
         have s_state: "same_state (Γ, (xs, s) # (as,sa) # ys) 
                           (map (λi. (Γ, (fst i, s) # snd i))
                                   (zip xs ?clist))"
              using hyp
              apply (simp add:hyp conjoin_def same_state_def  a1)              
              apply clarify
              apply(case_tac j) 
              by (simp add: xs_as_eq,simp add: xs_as_eq)
         have s_function: "same_functions (Γ, (xs, s) # (as,sa) # ys) 
                           (map (λi. (Γ, (fst i, s) # snd i))
                                   (zip xs ?clist))"
              using hyp conjoin_def same_functions_def a1 by fastforce
         have s_program: "same_program (Γ, (xs, s) # (as,sa) # ys) 
                           (map (λi. (Γ, (fst i, s) # snd i))
                                   (zip xs ?clist))"          
              using hyp
              apply (simp add:hyp conjoin_def same_program_def same_length_def a1)
              apply clarify
              apply(case_tac j)                
                apply(rule nth_equalityI) 
                apply(simp,simp)              
              by(rule nth_equalityI, simp add: hyp xs_as_eq, simp add:xs_as_eq)
         have s_compat:"compat_label (Γ, (xs, s) # (xs,sa) # ys) 
                           (map (λi. (Γ, (fst i, s) # snd i))
                                   (zip xs ?clist))" 
            using hyp a1 pe_ce
            apply (simp add:hyp conjoin_def compat_label_def)
            apply clarify
            apply(case_tac j,simp add: xs_as_eq)
               apply blast
              apply (simp add: xs_as_eq step_e.intros step_pe.intros)
             apply clarify
            apply(erule_tac x=nat in allE,erule impE,assumption)             
            apply(erule disjE,simp)
            apply clarify
            apply(rule_tac x=i in exI) 
            using hyp by (fastforce)+                            
        thus ?thesis using s_len s_program s_state s_function conjoin_def xs_as_eq
            by blast
     qed
     then have 
      "(∃clist.
                  length clist = length xs ∧
                  (Γ, (xs, s) #
                       a # ys) ∝ map (λi. (Γ, (fst i, s) # snd i))
                                   (zip xs clist) ∧
                  (∀i<length xs. (Γ, (xs ! i, s) # clist ! i) ∈ cptn))"
     using s1 s2 a_pair by blast
   } note s1=this

   {
     assume a1':"Γ⊢p (xs, s) → (as, sa)"
     then obtain i r where 
       inter_tran:"i < length xs ∧ Γ⊢c (xs ! i, s) → (r, sa) ∧ as = xs[i := r]" 
     using step_p_pair_elim_cases by metis     
     then have xs_as_eq_len: "length xs = length as" by simp
     from inter_tran 
      have s_states:"∃nsa. s=Normal nsa ∨ (s=sa ∧ (∀sa. (s≠Normal sa)))"
      using step_not_normal_s_eq_t by blast
     have as_xs:"∀i'<length as. (i'=i ∧ as!i'=r) ∨ (as!i'=xs!i')" 
       using xs_as_eq_len by (simp add: inter_tran nth_list_update)
     let ?clist="(map (λj. (as!j, sa)#(clist!j)) [0..<length xs]) [i:=((r, sa)#(clist!i))]"
     have s1:"length ?clist = length xs"
       by auto
     have s2:"(∀i'<length xs. (Γ, (xs ! i', s) # ?clist ! i') ∈ cptn)" 
        proof -
         {fix i'
          assume a1:"i' < length xs"          
          have "(Γ, (xs ! i', s) # ?clist ! i') ∈ cptn"
          proof (cases "i=i'")
            case True 
             thus ?thesis using inter_tran  hyp cptn.CptnComp               
              apply simp 
              by fastforce  
          next              
            case False            
            thus ?thesis using s_states inter_tran  False hyp cptn.CptnComp a1
              apply clarify
              apply simp
              apply(erule_tac x=i' in allE)
              apply (simp)
              apply(rule CptnEnv) 
              by (auto simp add: Env Env_n)
          qed
         } 
        thus ?thesis by fastforce
      qed
     then have s3:"(Γ, (xs, s) #
                       (as,sa) # ys) ∝ map (λi. (Γ, (fst i, s) # snd i))
                                   (zip xs ?clist)"
     proof -        
        from hyp have 
         len_list:"length clist = length as" by auto
        from hyp have same_len:"same_length (Γ, (as, sa) # ys)  
                      (map (λi. (Γ, (fst i, sa) # snd i)) (zip as clist))"
          using conjoin_def by auto        
        have s_len: "same_length (Γ, (xs, s) # (as,sa) # ys) 
                           (map (λi. (Γ, (fst i, s) # snd i))
                                   (zip xs ?clist))"  
          using 
            same_len  inter_tran  
            unfolding conjoin_def same_length_def
            apply clarify 
            apply(case_tac "i=ia")            
            by (auto simp add: len_list)            
        have s_state: "same_state (Γ, (xs, s) # (as,sa) # ys) 
                           (map (λi. (Γ, (fst i, s) # snd i))
                                   (zip xs ?clist))"
              using hyp inter_tran unfolding conjoin_def same_state_def
               apply clarify
               apply(case_tac j, simp, simp (no_asm_simp))
               apply(case_tac "i=ia",simp , simp )
              by (metis (no_types, hide_lams) as_xs nth_list_update_eq xs_as_eq_len)              
        have s_function: "same_functions (Γ, (xs, s) # (as,sa) # ys) 
                           (map (λi. (Γ, (fst i, s) # snd i))
                                   (zip xs ?clist))"
              using hyp conjoin_def same_functions_def a1 by fastforce
        have s_program: "same_program (Γ, (xs, s) # (as,sa) # ys) 
                           (map (λi. (Γ, (fst i, s) # snd i))
                                   (zip xs ?clist))"          
          using hyp inter_tran unfolding conjoin_def same_program_def
           apply clarify
           apply(case_tac j,simp)           
           apply(rule nth_equalityI,simp,simp)
           apply simp
           apply(rule nth_equalityI,simp,simp)
           apply(erule_tac x=nat and P="λj. H j ⟶ (fst (a j))=((b j))" for H a b in allE)
           apply(case_tac nat)
           apply clarify
           apply(case_tac "i=ia",simp,simp)
           apply clarify
           by(case_tac "i=ia",simp,simp)                   
        have s_compat:"compat_label (Γ, (xs, s) # (as,sa) # ys) 
                           (map (λi. (Γ, (fst i, s) # snd i))
                                   (zip xs ?clist))" 
        using inter_tran hyp s_states
        unfolding conjoin_def compat_label_def
         apply clarify
         apply(case_tac j)
          apply(rule conjI,simp)
           apply(erule ParComp,assumption)
           apply clarify
           apply(rule exI[where x=i],simp)
          apply clarify
          apply (rule snormal_enviroment,assumption)
         apply simp
         apply(erule_tac x=nat and P="λj. H j ⟶ (P j ∨ Q j)" for H P Q in allE,simp)
         apply (thin_tac "s = Normal nsa ∨ s = sa ∧ (∀sa. s ≠ Normal sa)")
        apply(erule disjE )
         apply clarify
         apply(rule_tac x=ia in exI,simp)
         apply(rule conjI)
          apply(case_tac "i=ia",simp,simp)
         apply clarify
         apply(case_tac "i=l",simp)
          apply(case_tac "l=ia",simp,simp)
          apply(erule_tac x=l in allE,erule impE,assumption,erule impE, assumption,simp)
         apply simp
         apply(erule_tac x=l in allE,erule impE,assumption,erule impE, assumption,simp)
        apply clarify
        apply (thin_tac " ∀ia<length xs. (Γ, (xs[i := r] ! ia, sa) # clist ! ia) ∈ cptn")
        apply(erule_tac x=ia and P="λj. H j ⟶ (P j)" for H P in allE, erule impE, assumption)
        by(case_tac "i=ia",simp,simp)               
        thus ?thesis using s_len s_program s_state s_function conjoin_def  
          by blast
     qed     
     then have "(∃clist.
                  length clist = length xs ∧
                  (Γ, (xs, s) #
                       a # ys) ∝ map (λi. (Γ, (fst i, s) # snd i))
                                   (zip xs clist) ∧
                  (∀i<length xs. (Γ, (xs ! i, s) # clist ! i) ∈ cptn))"
     using s1 s2 a_pair by blast
   } 
   then have 
      "(∃clist.
                  length clist = length xs ∧
                  (Γ, (xs, s) #
                       a # ys) ∝ map (λi. (Γ, (fst i, s) # snd i))
                                   (zip xs clist) ∧
                  (∀i<length xs. (Γ, (xs ! i, s) # clist ! i) ∈ cptn))"
      using s1 par_cptn_dest by fastforce
  }  
  thus ?case by auto
qed  

lemma one_iff_aux_if:"xs≠[] ⟹ (∀ys. ((Γ,((xs, s)#ys)) ∈ par_cptn) = 
 (∃clist. length clist= length xs ∧ 
 ((Γ,(xs, s)#ys) ∝ map (λi. (Γ,(fst i,s)#(snd i))) (zip xs clist)) ∧ 
 (∀i<length xs. (Γ,(xs!i,s)#(clist!i)) ∈ cptn))) ⟹
 (par_cp Γ (xs) s = {(Γ1,c). ∃clist. (length clist)=(length xs) ∧
 (∀i<length clist. clist!i ∈ cp Γ (xs!i) s) ∧ (Γ,c) ∝ clist ∧ Γ1=Γ})"
proof
  assume a1:"xs≠[]"
  assume a2:"∀ys. ((Γ, (xs, s) # ys) ∈ par_cptn) =
         (∃clist.
             length clist = length xs ∧
             (Γ,
              (xs, s) #
              ys) ∝ map (λi. (Γ, (fst i, s) # snd i))
                      (zip xs clist) ∧
             (∀i<length xs.
                 (Γ, (xs ! i, s) # clist ! i) ∈ cptn))"         
   show "par_cp Γ xs s ⊆ 
             {(Γ1, c). ∃clist.
               length clist = length xs ∧
               (∀i<length clist. clist ! i ∈ cp Γ (xs ! i) s) ∧
               (Γ, c) ∝ clist ∧ Γ1 = Γ}"
   proof-{     
     fix x
     let ?show = "x∈ {(Γ1, c). ∃clist.
       length clist = length xs ∧
       (∀i<length clist. clist ! i ∈ cp Γ (xs ! i) s) ∧
        (Γ, c) ∝ clist ∧ Γ1 = Γ}"
     assume a3:"x∈par_cp Γ xs s"   
     then obtain y where x_pair: "x=(Γ,y)"
       unfolding par_cp_def by auto       
     have ?show          
     proof (cases y)
        case Nil then 
         show ?show using a1 a2 a3 x_pair
          unfolding par_cp_def cp_def
          by (force elim:par_cptn.cases)
     next 
        case (Cons a list) then
          show ?show using a1 a2 a3 x_pair
          unfolding par_cp_def cp_def         
          by(auto, rule_tac x="map (λi. (Γ,(fst i, s) # snd i)) (zip xs clist)" in exI,simp)
     qed
   } thus ?thesis using a1 a2 by auto 
   qed
   {   
   show "{(Γ1, c). ∃clist.
          length clist = length xs ∧
          (∀i<length clist. clist ! i ∈ cp Γ (xs ! i) s) ∧
          (Γ, c) ∝ clist ∧ Γ1 = Γ} ⊆ par_cp Γ xs s" using a1 a2 
   proof-
     { 
     fix x
     assume a3:"x∈{(Γ1, c). ∃clist.
          length clist = length xs ∧
          (∀i<length clist. clist ! i ∈ cp Γ (xs ! i) s) ∧
          (Γ, c) ∝ clist ∧ Γ1 = Γ}"
     then obtain c where x_pair: "x=(Γ,c)"  by auto
     then obtain clist where 
      props:"length clist = length xs ∧
           (∀i<length clist. clist ! i ∈ cp Γ (xs ! i) s) ∧
           (Γ, c) ∝ clist " using a3 by auto
     then have "x∈par_cp Γ xs s"
       proof (cases c)
         case Nil 
         have clist_0: 
           "clist ! 0 ∈ cp Γ (xs ! 0) s" using props a1 
         by auto
         thus "x∈par_cp Γ xs s"  
           using a1 a2 props Nil x_pair
         unfolding cp_def conjoin_def same_length_def 
         apply clarify                  
         by(erule cptn.cases,fastforce,fastforce,fastforce)
       next
         case (Cons a ys) 
         then obtain a1 a2 where a_pair: "a=(a1,a2)" 
           using props by fastforce 
         from a2 have 
               a2:"(((Γ, (xs, s) # ys) ∈ par_cptn) =
                   (∃clist. length clist = length xs ∧
                   (Γ, (xs, s) # ys) ∝ map (λi. (Γ, (fst i, s) # snd i)) (zip xs clist) ∧
                   (∀i<length xs. (Γ, (xs ! i, s) # clist ! i) ∈ cptn)))" by auto
         have a2_s:"a2=s" using a1 props a_pair Cons
           unfolding  conjoin_def   same_state_def  cp_def         
           by force
         have a1_xs:"a1 = xs"
           using  props a_pair Cons 
           unfolding par_cp_def conjoin_def  same_program_def cp_def           
           apply clarify
           apply(erule_tac x=0 and P="λj. H j ⟶ (fst (s j))=((t j))" for H s t in allE)                      
           by(rule nth_equalityI,auto)   
         then have conjoin_clist_xs:"(Γ, (xs,s)#ys) ∝ clist"     
           using a1  props a_pair Cons a1_xs a2_s by auto
         also then have "clist = map (λi. (Γ,(fst i,s)#(snd i))) (zip xs ((map (λx. tl (snd x))) clist))"             
           using clist_map_zip a1  by fastforce
         ultimately have conjoin_map:"(Γ, (xs, s) # ys) ∝ map (λi. (Γ, (fst i, s) # snd i)) (zip xs ((map (λx. tl (snd x))) clist))"
           using props x_pair Cons a_pair a1_xs a2_s by auto    
         have "⋀n. ¬ n < length xs ∨ clist ! n ∈ {(f, ps). ps ! 0 = (xs ! n, a2) ∧ (Γ, ps) ∈ cptn ∧ f = Γ}"
               using a1_xs a2_s props cp_def by fastforce
         then have clist_cptn:"(∀i<length clist. (fst (clist!i) = Γ) ∧ 
                                 (Γ, snd (clist!i)) ∈ cptn ∧
                                 (snd (clist!i))!0 = (xs!i,s))"
         using a1_xs a2_s props by fastforce         
                       
          {fix i
          assume a4: "i<length xs"     
          then have clist_i_cptn:"(fst (clist!i) = Γ) ∧ 
                     (Γ, snd (clist!i)) ∈ cptn ∧
                     (snd (clist!i))!0 = (xs!i,s)"
           using props clist_cptn by fastforce 
          from a4 props have a4':"i<length clist" by auto
          have lengz:"length (snd (clist!i))>0" 
            using conjoin_clist_xs a4'
            unfolding  conjoin_def same_length_def 
           by auto
          then have clist_hd_tl:"snd (clist!i) =  hd (snd (clist!i)) # tl (snd (clist ! i))"
            by auto        
          also have " hd (snd (clist!i)) =  (snd (clist!i))!0"
            using a4' lengz by (simp add: hd_conv_nth)
          ultimately have clist_i_tl:"snd (clist!i) =  (xs!i,s) # tl (snd (clist ! i))"
            using clist_i_cptn by fastforce
          also have "tl (snd (clist ! i)) = map (λx. tl (snd x)) clist!i"
            using nth_map a4' 
          by auto
          ultimately have snd_clist:"snd (clist!i) =  (xs ! i, s) # map (λx. tl (snd x)) clist ! i"
            by auto
          also have "(clist!i) = (fst (clist!i),snd (clist!i))"
            by auto
          ultimately have "(clist!i) =(Γ, (xs ! i, s) # map (λx. tl (snd x)) clist ! i)"
           using clist_i_cptn by auto
          then have "(Γ, (xs ! i, s) # map (λx. tl (snd x)) clist ! i) ∈ cptn" 
             using clist_i_cptn by auto
          } 
          then have clist_in_cptn:"(∀i<length xs. (Γ, (xs ! i, s) # ((map (λx. tl (snd x))) clist) ! i) ∈ cptn)"
          by auto
         have same_length_clist_xs:"length ((map (λx. tl (snd x))) clist)  = length xs"
           using props by auto
         then have "(∃clist. length clist = length xs ∧
                        (Γ, (xs, s) # ys) ∝ map (λi. (Γ, (fst i, s) # snd i)) (zip xs clist) ∧
                        (∀i<length xs. (Γ, (xs ! i, s) # clist ! i) ∈ cptn))"
         using  a1  props x_pair a_pair Cons a1_xs a2_s conjoin_clist_xs clist_in_cptn
            conjoin_map clist_map by blast         
         then have "(Γ, c) ∈ par_cptn" using  a1 a2  props x_pair a_pair Cons a1_xs a2_s
         unfolding par_cp_def by simp          
         thus "x∈par_cp Γ xs s"  
           using a1 a2  props x_pair a_pair Cons a1_xs a2_s
         unfolding par_cp_def conjoin_def  same_length_def same_program_def same_state_def same_functions_def compat_label_def 
           by simp          
       qed
     }
     thus ?thesis using a1 a2 by auto  
   qed
  } 
qed



lemma one_iff_aux_only_if:"xs≠[] ⟹  
 (par_cp Γ (xs) s = {(Γ1,c). ∃clist. (length clist)=(length xs) ∧
 (∀i<length clist. clist!i ∈ cp Γ (xs!i) s) ∧ (Γ,c) ∝ clist ∧ Γ1=Γ}) ⟹
(∀ys. ((Γ,((xs, s)#ys)) ∈ par_cptn) = 
 (∃clist. length clist= length xs ∧ 
 ((Γ,(xs, s)#ys) ∝ map (λi. (Γ,(fst i,s)#(snd i))) (zip xs clist)) ∧ 
 (∀i<length xs. (Γ,(xs!i,s)#(clist!i)) ∈ cptn)))"
proof
  fix ys
  assume a1: "xs≠[]"
  assume a2: "par_cp Γ xs s =
          {(Γ1, c).
           ∃clist.
              length clist = length xs ∧
              (∀i<length clist.
                  clist ! i ∈ cp Γ (xs ! i) s) ∧
              (Γ, c) ∝ clist ∧ Γ1 = Γ}"
  from a1 a2 show
  "((Γ, (xs, s) # ys) ∈ par_cptn) =
          (∃clist.
              length clist = length xs ∧
              (Γ,
               (xs, s) #
               ys) ∝ map (λi. (Γ, (fst i, s) # snd i))
                       (zip xs clist) ∧
              (∀i<length xs.
                  (Γ, (xs ! i, s) # clist ! i) ∈ cptn))"
  proof auto
    {assume a3:"(Γ, (xs, s) # ys) ∈ par_cptn"
     then show "∃clist.
       length clist = length xs ∧
       (Γ,
        (xs, s) #
        ys) ∝ map (λi. (Γ, (fst i, s) # snd i))
                (zip xs clist) ∧
       (∀i<length xs. (Γ, (xs ! i, s) # clist ! i) ∈ cptn)"
       using a1 a2 by (simp add: aux_onlyif)
    }
    {fix clist ::"(('a, 'b, 'c, 'd) LanguageCon.com ×
             ('a, 'c) xstate) list list"
    assume a3: "length clist = length xs"
    assume a4:"(Γ, (xs, s) # ys) ∝ 
               map (λi. (Γ, (fst i, s) # snd i))
                (zip xs clist)"
    assume a5: "∀i<length xs. (Γ, (xs ! i, s) # clist ! i)
                     ∈ cptn"
    show "(Γ, (xs, s) # ys) ∈ par_cptn" 
    using a3 a4 a5 using aux_if by blast 
    }
  qed
qed

lemma one_iff_aux: "xs≠[] ⟹ (∀ys. ((Γ,((xs, s)#ys)) ∈ par_cptn) = 
 (∃clist. length clist= length xs ∧ 
 ((Γ,(xs, s)#ys) ∝ map (λi. (Γ,(fst i,s)#(snd i))) (zip xs clist)) ∧ 
 (∀i<length xs. (Γ,(xs!i,s)#(clist!i)) ∈ cptn))) = 
 (par_cp Γ (xs) s = {(Γ1,c). ∃clist. (length clist)=(length xs) ∧
 (∀i<length clist. clist!i ∈ cp Γ (xs!i) s) ∧ (Γ,c) ∝ clist ∧ Γ1=Γ})"
proof 
  assume a1:"xs≠[]"
  {assume a2:"(∀ys. ((Γ,((xs, s)#ys)) ∈ par_cptn) = 
   (∃clist. length clist= length xs ∧ 
   ((Γ,(xs, s)#ys) ∝ map (λi. (Γ,(fst i,s)#(snd i))) (zip xs clist)) ∧ 
   (∀i<length xs. (Γ,(xs!i,s)#(clist!i)) ∈ cptn)))"
    then show "(par_cp Γ (xs) s = {(Γ1,c). ∃clist. (length clist)=(length xs) ∧
   (∀i<length clist. clist!i ∈ cp Γ (xs!i) s) ∧ (Γ,c) ∝ clist ∧ Γ1=Γ})"
    by (auto simp add: a1 a2 one_iff_aux_if)
  }
  {assume a2:"(par_cp Γ (xs) s = {(Γ1,c). ∃clist. (length clist)=(length xs) ∧
   (∀i<length clist. clist!i ∈ cp Γ (xs!i) s) ∧ (Γ,c) ∝ clist ∧ Γ1=Γ})"    
    then show "(∀ys. ((Γ,((xs, s)#ys)) ∈ par_cptn) = 
   (∃clist. length clist= length xs ∧ 
   ((Γ,(xs, s)#ys) ∝ map (λi. (Γ,(fst i,s)#(snd i))) (zip xs clist)) ∧ 
   (∀i<length xs. (Γ,(xs!i,s)#(clist!i)) ∈ cptn)))"
   by (auto simp add: a1 a2 one_iff_aux_only_if)
  }
qed
  


theorem one: 
"xs≠[] ⟹  
 par_cp Γ xs s = 
    {(Γ1,c). ∃clist. (length clist)=(length xs) ∧ 
             (∀i<length clist. (clist!i) ∈ cp Γ (xs!i) s) ∧ 
             (Γ,c) ∝ clist ∧ Γ1=Γ}"

apply(frule one_iff_aux)
apply(drule sym)
apply(erule iffD2)
apply clarify
apply(rule iffI) 
 apply(erule aux_onlyif)
apply clarify
apply(force intro:aux_if)
done

(************************************************************************ *)
(* subsection {* Equivalence between Small-Step and Big-Step Semantics *} *)
(* ************************************************************************ *)
 
(* 

?Γ⊢c (LanguageCon.com.Seq ?c1.0 ?c2.0, ?s) →
       (LanguageCon.com.Seq ?c1' ?c2.0,
        ?s') ⟹
(?Γ⊢c (?c1.0, ?s) → (?c1', ?s') ⟹ ?P) ⟹
?P

lemma 
   assumes 
           step_a: "Γ⊢c(Await b c, Normal s) → (t,u)"
   shows step_await_step_c:"(Γ¬a)⊢(c, Normal s) →* (sequential t,u)" 
using step_a
proof cases
  fix t1
  assume
      "(t, u) = (Skip, t1)" "s ∈ b" "Γ⊢ ⟨c,Normal s⟩ ⇒ t1" "∀t'. t1 ≠ Abrupt t'"
  thus ?thesis 
   by (cases u) 
   (auto intro: exec_impl_steps_Fault exec_impl_steps_Normal exec_impl_steps_Stuck)
next
  fix t1
  assume "(t, u) = (Throw, Normal t1)" "s ∈ b" "Γ⊢ ⟨c,Normal s⟩ ⇒ Abrupt t1"
  thus ?thesis by (simp add: exec_impl_steps_Normal_Abrupt)
qed

lemma 
   assumes (* exec: "Γ⊢⟨c,Normal s⟩ ⇒ t" and
           b: "s ∈ b" and *)
           step_a: "Γ⊢c(Await b c, Normal s) → u"
   shows step_await_final1:"final u"
using step_a 
proof cases
  case  (1 t) thus "final u"  by (simp add: final_def)
next
  case (2 t)
  thus "final u" by (simp add: exec_impl_steps_Normal_Abrupt final_def)
qed

lemma step_Abrupt_end: 
  assumes step: "Γ⊢c (c1, s) → (c1', s')"
  shows "s'=Abrupt x ⟹ s=Abrupt x"
using step
by induct auto


lemma step_Stuck_end: 
  assumes step: "Γ⊢c (c1, s) → (c1', s')"
  shows "s'=Stuck ⟹ 
          s=Stuck ∨ 
          (∃r x. redex c1 = Spec r ∧ s=Normal x ∧ (∀t. (x,t)∉r)) ∨ 
          (∃p x. redex c1=  Call p ∧ s=Normal x ∧ Γ p = None) ∨
          (∃b c x.  redex c1 = Await b c ∧ s=Normal x ∧ x ∈ b ∧ Γ⊢⟨c,s⟩⇒s')"
using step
by induct auto

lemma step_Fault_end: 
  assumes step: "Γ⊢c (c1, s) → (c1', s')"
  shows "s'=Fault f ⟹ 
          s=Fault f ∨ 
          (∃g c x.  redex c1 = Guard f g c ∧ s=Normal x ∧ x ∉ g) ∨
          (∃b c1 x.  redex c1 = Await b c1 ∧ s=Normal x ∧ x ∈ b ∧ Γ⊢⟨c1,s⟩⇒s')"
using step 
by induct auto



(* ************************************************************************ *)
subsection {* Infinite Computations: @{text "Γ⊢(c, s) → …(∞)"}*}
(* ************************************************************************ *)

definition inf_c:: "('s,'p,'f,'e) body ⇒ ('s,'p,'f,'e) config ⇒ bool"
 ("_⊢c _ → …'(∞')" [60,80] 100) where
"Γ⊢c cfg → …(∞) ≡ (∃f. f (0::nat) = cfg ∧ (∀i. Γ⊢cf i → f (i+1)))" 

lemma not_infI: "⟦⋀f. ⟦f 0 = cfg; ⋀i. Γ⊢cf i → f (Suc i)⟧ ⟹ False⟧  
                ⟹ ¬Γ⊢c cfg → …(∞)"
  by (auto simp add: inf_c_def)

(* ************************************************************************ *)
subsection {* Equivalence between Termination and the Absence of Infinite Computations*}
(* ************************************************************************ *)


lemma step_preserves_termination: 
  assumes step: "Γ⊢(c,s) → (c',s')"
  shows "Γ⊢c↓s ⟹ Γ⊢c'↓s'"  
using step
proof (induct)
  case Basic thus ?case by (fastforce intro: terminates.intros)
next
  case Spec thus ?case by (fastforce intro: terminates.intros)
next
  case SpecStuck thus ?case by (fastforce intro: terminates.intros)
next
  case Guard thus ?case 
    by (fastforce intro: terminates.intros elim: terminates_Normal_elim_cases)
next
  case GuardFault thus ?case by (fastforce intro: terminates.intros)
next
  case (Seq c1 s c1' s' c2) thus ?case
    apply (cases s)
    apply     (cases s')
    apply         (fastforce intro: terminates.intros step_extend 
                    elim: terminates_Normal_elim_cases)
    apply (fastforce intro: terminates.intros dest: step_Abrupt_prop 
      step_Fault_prop step_Stuck_prop)+
    done
next
  case (SeqSkip c2 s) 
  thus ?case 
    apply (cases s)
    apply (fastforce intro: terminates.intros exec.intros
            elim: terminates_Normal_elim_cases )+
    done
next
  case (SeqThrow c2 s) 
  thus ?case 
    by (fastforce intro: terminates.intros exec.intros
            elim: terminates_Normal_elim_cases )
next
  case CondTrue 
  thus ?case 
    by (fastforce intro: terminates.intros exec.intros
            elim: terminates_Normal_elim_cases )
next
  case CondFalse 
  thus ?case 
    by (fastforce intro: terminates.intros 
            elim: terminates_Normal_elim_cases )
next
  case WhileTrue
  thus ?case 
    by (fastforce intro: terminates.intros 
            elim: terminates_Normal_elim_cases )
next
  case WhileFalse 
  thus ?case 
    by (fastforce intro: terminates.intros 
            elim: terminates_Normal_elim_cases )
next
  case Call 
  thus ?case 
    by (fastforce intro: terminates.intros 
            elim: terminates_Normal_elim_cases )
next
  case CallUndefined
  thus ?case 
    by (fastforce intro: terminates.intros 
            elim: terminates_Normal_elim_cases )
next
  case DynCom
  thus ?case 
    by (fastforce intro: terminates.intros 
            elim: terminates_Normal_elim_cases )
next
  case (Catch c1 s c1' s' c2) thus ?case
    apply (cases s)
    apply     (cases s')
    apply         (fastforce intro: terminates.intros step_extend 
                    elim: terminates_Normal_elim_cases)
    apply (fastforce intro: terminates.intros dest: step_Abrupt_prop 
      step_Fault_prop step_Stuck_prop)+
    done
next
  case CatchThrow
  thus ?case 
   by (fastforce intro: terminates.intros exec.intros
            elim: terminates_Normal_elim_cases )
next
  case (CatchSkip c2 s) 
  thus ?case 
    by (cases s) (fastforce intro: terminates.intros)+
next
  case FaultProp thus ?case by (fastforce intro: terminates.intros)
next
  case StuckProp thus ?case by (fastforce intro: terminates.intros)
next
  case AbruptProp thus ?case by (fastforce intro: terminates.intros)
next 
  case Await thus ?case using terminates_Skip' by blast 
next 
  case AwaitAbrupt thus ?case by (fastforce intro: terminates.intros)
qed

lemma steps_preserves_termination: 
  assumes steps: "Γ⊢(c,s) →* (c',s')"
  shows "Γ⊢c↓s ⟹ Γ⊢c'↓s'"
using steps
proof (induct rule: rtranclp_induct2 [consumes 1, case_names Refl Trans])
  case Refl thus ?case  .
next
  case Trans
  thus ?case
    by (blast dest: step_preserves_termination)
qed

ML {*
  ML_Thms.bind_thm ("tranclp_induct2", Split_Rule.split_rule @{context}
    (Rule_Insts.read_instantiate @{context}
      [((("a", 0), Position.none), "(aa,ab)"), ((("b", 0), Position.none), "(ba,bb)")] []
      @{thm tranclp_induct}));
*}

thm tranclp_induct2 tranclp_induct

lemma steps_preserves_termination': 
  assumes steps: "Γ⊢(c,s) →+ (c',s')"
  shows "Γ⊢c↓s ⟹ Γ⊢c'↓s'"
using steps
proof (induct rule: tranclp_induct2 [consumes 1, case_names Step Trans])
  case Step thus ?case by (blast intro: step_preserves_termination)
next
  case Trans
  thus ?case
    by (blast dest: step_preserves_termination)
qed



definition head_com:: "('s,'p,'f,'e) com ⇒ ('s,'p,'f,'e) com"
where
"head_com c =
  (case c of
     Seq c1 c2 ⇒ c1
   | Catch c1 c2 ⇒ c1
   | _ ⇒ c)"

  
definition head:: "('s,'p,'f,'e) config ⇒ ('s,'p,'f,'e) config"
  where "head cfg = (head_com (fst cfg), snd cfg)"

lemma le_Suc_cases: "⟦⋀i. ⟦i < k⟧ ⟹ P i; P k⟧ ⟹ ∀i<(Suc k). P i"
  apply clarify
  apply (case_tac "i=k")
  apply auto
  done

lemma redex_Seq_False: "⋀c' c''. (redex c = Seq c'' c') = False"
  by (induct c) auto

lemma redex_Catch_False: "⋀c' c''. (redex c = Catch c'' c') = False"
  by (induct c) auto


lemma infinite_computation_extract_head_Seq:
  assumes inf_comp: "∀i::nat. Γ⊢f i → f (i+1)"
  assumes f_0: "f 0 = (Seq c1 c2,s)"
  assumes not_fin: "∀i<k. ¬ final (head (f i))"
  shows "∀i<k. (∃c' s'. f (i + 1) = (Seq c' c2, s')) ∧  
               Γ⊢head (f i) → head (f (i+1))"
        (is "∀i<k. ?P i")
using not_fin
proof (induct k)
  case 0
  show ?case by simp
next
  case (Suc k)
  have not_fin_Suc: 
    "∀i<Suc k. ¬ final (head (f i))" by fact
  from this[rule_format] have not_fin_k: 
    "∀i<k. ¬ final (head (f i))" 
    apply clarify
    apply (subgoal_tac "i < Suc k")
    apply blast
    apply simp
    done

  from Suc.hyps [OF this]
  have hyp: "∀i<k. (∃c' s'. f (i + 1) = (Seq c' c2, s')) ∧ 
                   Γ⊢ head (f i) → head (f (i + 1))".
  show ?case
  proof (rule le_Suc_cases)
    fix i 
    assume "i < k"
    then show "?P i"
      by (rule hyp [rule_format])
  next
    show "?P k"
    proof -
      from hyp [rule_format, of "k - 1"] f_0
      obtain c' fs' L' s' where  f_k: "f k = (Seq c' c2, s')"
        by (cases k) auto
      from inf_comp [rule_format, of k] f_k
      have "Γ⊢(Seq c' c2, s') → f (k + 1)"
        by simp
      moreover
      from not_fin_Suc [rule_format, of k] f_k
      have "¬ final (c',s')"
        by (simp add: final_def head_def head_com_def)
      ultimately
      obtain c'' s'' where
         "Γ⊢(c', s') → (c'', s'')" and
         "f (k + 1) = (Seq c'' c2, s'')"
        by cases (auto simp add: redex_Seq_False final_def)
      with f_k
      show ?thesis
        by (simp add: head_def head_com_def)
    qed
  qed
qed

lemma infinite_computation_extract_head_Catch:
  assumes inf_comp: "∀i::nat. Γ⊢f i → f (i+1)"
  assumes f_0: "f 0 = (Catch c1 c2,s)"
  assumes not_fin: "∀i<k. ¬ final (head (f i))"
  shows "∀i<k. (∃c' s'. f (i + 1) = (Catch c' c2, s')) ∧  
               Γ⊢head (f i) → head (f (i+1))"
        (is "∀i<k. ?P i")
using not_fin
proof (induct k)
  case 0
  show ?case by simp
next
  case (Suc k)
  have not_fin_Suc: 
    "∀i<Suc k. ¬ final (head (f i))" by fact
  from this[rule_format] have not_fin_k: 
    "∀i<k. ¬ final (head (f i))" 
    apply clarify
    apply (subgoal_tac "i < Suc k")
    apply blast
    apply simp
    done

  from Suc.hyps [OF this]
  have hyp: "∀i<k. (∃c' s'. f (i + 1) = (Catch c' c2, s')) ∧ 
                   Γ⊢ head (f i) → head (f (i + 1))".
  show ?case
  proof (rule le_Suc_cases)
    fix i 
    assume "i < k"
    then show "?P i"
      by (rule hyp [rule_format])
  next
    show "?P k"
    proof -
      from hyp [rule_format, of "k - 1"] f_0
      obtain c' fs' L' s' where  f_k: "f k = (Catch c' c2, s')"
        by (cases k) auto
      from inf_comp [rule_format, of k] f_k
      have "Γ⊢(Catch c' c2, s') → f (k + 1)"
        by simp
      moreover
      from not_fin_Suc [rule_format, of k] f_k
      have "¬ final (c',s')"
        by (simp add: final_def head_def head_com_def)
      ultimately
      obtain c'' s'' where
         "Γ⊢(c', s') → (c'', s'')" and
         "f (k + 1) = (Catch c'' c2, s'')"
        by cases (auto simp add: redex_Catch_False final_def)+
      with f_k
      show ?thesis
        by (simp add: head_def head_com_def)
    qed
  qed
qed

lemma no_inf_Throw: "¬ Γ⊢(Throw,s) → …(∞)"
proof 
  assume "Γ⊢ (Throw, s) → …(∞)"
  then obtain f where
    step [rule_format]: "∀i::nat. Γ⊢f i → f (i+1)" and
    f_0: "f 0 = (Throw, s)"
    by (auto simp add: inf_def)
  from step [of 0, simplified f_0] step [of 1]
  show False
    by cases (auto elim: step_elim_cases)
qed

lemma split_inf_Seq: 
  assumes inf_comp: "Γ⊢(Seq c1 c2,s) → …(∞)"
  shows "Γ⊢(c1,s) → …(∞) ∨ 
         (∃s'. Γ⊢(c1,s) →* (Skip,s') ∧ Γ⊢(c2,s') → …(∞))"
proof -
  from inf_comp obtain f where
    step: "∀i::nat. Γ⊢f i → f (i+1)" and
    f_0: "f 0 = (Seq c1 c2, s)"
    by (auto simp add: inf_def)
  from f_0 have head_f_0: "head (f 0) = (c1,s)" 
    by (simp add: head_def head_com_def)
  show ?thesis
  proof (cases "∃i. final (head (f i))")
    case True
    def k≡"(LEAST i. final (head (f i)))"
    have less_k: "∀i<k. ¬ final (head (f i))"
      apply (intro allI impI)
      apply (unfold k_def)
      apply (drule not_less_Least)
      apply auto
      done
    from infinite_computation_extract_head_Seq [OF step f_0 this]
    obtain step_head: "∀i<k. Γ⊢ head (f i) → head (f (i + 1))" and
           conf: "∀i<k. (∃c' s'. f (i + 1) = (Seq c' c2, s'))"
      by blast 
    from True
    have final_f_k: "final (head (f k))"
      apply -
      apply (erule exE)
      apply (drule LeastI)
      apply (simp add: k_def)
      done
    moreover
    from f_0 conf [rule_format, of "k - 1"]
    obtain c' s' where f_k: "f k = (Seq c' c2,s')"
      by (cases k) auto
    moreover
    from step_head have steps_head: "Γ⊢head (f 0) →* head (f k)"
    proof (induct k)
      case 0 thus ?case by simp
    next
      case (Suc m)
      have step: "∀i<Suc m. Γ⊢ head (f i) → head (f (i + 1))" by fact
      hence "∀i<m. Γ⊢ head (f i) → head (f (i + 1))"
        by auto
      hence "Γ⊢ head (f 0) →*  head (f m)"
        by (rule Suc.hyps)
      also from step [rule_format, of m]
      have "Γ⊢ head (f m) → head (f (m + 1))" by simp
      finally show ?case by simp
    qed
    {
      assume f_k: "f k = (Seq Skip c2, s')"
      with steps_head
      have "Γ⊢(c1,s) →* (Skip,s')"
        using head_f_0
        by (simp add: head_def head_com_def)
      moreover
      from step [rule_format, of k] f_k
      obtain "Γ⊢(Seq Skip c2,s') → (c2,s')" and
        f_Suc_k: "f (k + 1) = (c2,s')"
        by (fastforce elim: step.cases intro: step.intros)
      def g≡"λi. f (i + (k + 1))"
      from f_Suc_k
      have g_0: "g 0 = (c2,s')"
        by (simp add: g_def)
      from step
      have "∀i. Γ⊢g i → g (i + 1)"
        by (simp add: g_def)
      with g_0 have "Γ⊢(c2,s') → …(∞)"
        by (auto simp add: inf_def)
      ultimately
      have ?thesis
        by auto
    }
    moreover
    {
      fix x
      assume s': "s'=Normal x" and f_k: "f k = (Seq Throw c2, s')"
      from step [rule_format, of k] f_k s'
      obtain "Γ⊢(Seq Throw c2,s') → (Throw,s')" and
        f_Suc_k: "f (k + 1) = (Throw,s')"
        by (fastforce elim: step_elim_cases intro: step.intros)
      def g≡"λi. f (i + (k + 1))"
      from f_Suc_k
      have g_0: "g 0 = (Throw,s')"
        by (simp add: g_def)
      from step
      have "∀i. Γ⊢g i → g (i + 1)"
        by (simp add: g_def)
      with g_0 have "Γ⊢(Throw,s') → …(∞)"
        by (auto simp add: inf_def)
      with no_inf_Throw
      have ?thesis
        by auto
    }
    ultimately 
    show ?thesis
      by (auto simp add: final_def head_def head_com_def)
  next
    case False
    then have not_fin: "∀i. ¬ final (head (f i))"
      by blast
    have "∀i. Γ⊢head (f i) → head (f (i + 1))"
    proof 
      fix k
      from not_fin 
      have "∀i<(Suc k). ¬ final (head (f i))"
        by simp
      
      from infinite_computation_extract_head_Seq [OF step f_0 this ]
      show "Γ⊢ head (f k) → head (f (k + 1))" by simp
    qed
    with head_f_0 have "Γ⊢(c1,s) → …(∞)"
      by (auto simp add: inf_def)
    thus ?thesis
      by simp
  qed
qed

lemma split_inf_Catch: 
  assumes inf_comp: "Γ⊢(Catch c1 c2,s) → …(∞)"
  shows "Γ⊢(c1,s) → …(∞) ∨ 
         (∃s'. Γ⊢(c1,s) →* (Throw,Normal s') ∧ Γ⊢(c2,Normal s') → …(∞))"
proof -
  from inf_comp obtain f where
    step: "∀i::nat. Γ⊢f i → f (i+1)" and
    f_0: "f 0 = (Catch c1 c2, s)"
    by (auto simp add: inf_def)
  from f_0 have head_f_0: "head (f 0) = (c1,s)" 
    by (simp add: head_def head_com_def)
  show ?thesis
  proof (cases "∃i. final (head (f i))")
    case True
    def k≡"(LEAST i. final (head (f i)))"
    have less_k: "∀i<k. ¬ final (head (f i))"
      apply (intro allI impI)
      apply (unfold k_def)
      apply (drule not_less_Least)
      apply auto
      done
    from infinite_computation_extract_head_Catch [OF step f_0 this]
    obtain step_head: "∀i<k. Γ⊢ head (f i) → head (f (i + 1))" and
           conf: "∀i<k. (∃c' s'. f (i + 1) = (Catch c' c2, s'))"
      by blast 
    from True
    have final_f_k: "final (head (f k))"
      apply -
      apply (erule exE)
      apply (drule LeastI)
      apply (simp add: k_def)
      done
    moreover
    from f_0 conf [rule_format, of "k - 1"]
    obtain c' s' where f_k: "f k = (Catch c' c2,s')"
      by (cases k) auto
    moreover
    from step_head have steps_head: "Γ⊢head (f 0) →* head (f k)"
    proof (induct k)
      case 0 thus ?case by simp
    next
      case (Suc m)
      have step: "∀i<Suc m. Γ⊢ head (f i) → head (f (i + 1))" by fact
      hence "∀i<m. Γ⊢ head (f i) → head (f (i + 1))"
        by auto
      hence "Γ⊢ head (f 0) →*  head (f m)"
        by (rule Suc.hyps)
      also from step [rule_format, of m]
      have "Γ⊢ head (f m) → head (f (m + 1))" by simp
      finally show ?case by simp
    qed
    {
      assume f_k: "f k = (Catch Skip c2, s')"
      with steps_head
      have "Γ⊢(c1,s) →* (Skip,s')"
        using head_f_0
        by (simp add: head_def head_com_def)
      moreover
      from step [rule_format, of k] f_k
      obtain "Γ⊢(Catch Skip c2,s') → (Skip,s')" and
        f_Suc_k: "f (k + 1) = (Skip,s')"
        by (fastforce elim: step.cases intro: step.intros)
      from step [rule_format, of "k+1", simplified f_Suc_k]
      have ?thesis
        by (rule no_step_final') (auto simp add: final_def)
    }
    moreover
    {
      fix x
      assume s': "s'=Normal x" and f_k: "f k = (Catch Throw c2, s')"
      with steps_head
      have "Γ⊢(c1,s) →* (Throw,s')"
        using head_f_0
        by (simp add: head_def head_com_def)
      moreover
      from step [rule_format, of k] f_k s'
      obtain "Γ⊢(Catch Throw c2,s') → (c2,s')" and
        f_Suc_k: "f (k + 1) = (c2,s')"
        by (fastforce elim: step_elim_cases intro: step.intros)
      def g≡"λi. f (i + (k + 1))"
      from f_Suc_k
      have g_0: "g 0 = (c2,s')"
        by (simp add: g_def)
      from step
      have "∀i. Γ⊢g i → g (i + 1)"
        by (simp add: g_def)
      with g_0 have "Γ⊢(c2,s') → …(∞)"
        by (auto simp add: inf_def)
      ultimately
      have ?thesis
        using s'
        by auto
    }
    ultimately 
    show ?thesis
      by (auto simp add: final_def head_def head_com_def)
  next
    case False
    then have not_fin: "∀i. ¬ final (head (f i))"
      by blast
    have "∀i. Γ⊢head (f i) → head (f (i + 1))"
    proof 
      fix k
      from not_fin 
      have "∀i<(Suc k). ¬ final (head (f i))"
        by simp
      
      from infinite_computation_extract_head_Catch [OF step f_0 this ]
      show "Γ⊢ head (f k) → head (f (k + 1))" by simp
    qed
    with head_f_0 have "Γ⊢(c1,s) → …(∞)"
      by (auto simp add: inf_def)
    thus ?thesis
      by simp
  qed
qed

lemma Skip_no_step: "Γ⊢(Skip,s) → cfg ⟹ P"
  apply (erule no_step_final')
  apply (simp add: final_def)
  done

lemma not_inf_Stuck: "¬ Γ⊢(c,Stuck) → …(∞)"
proof (induct c)
  case Skip
  show ?case
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Skip, Stuck)" 
    from f_step [of 0] f_0
    show False
      by (auto elim: Skip_no_step)
  qed
next
  case (Basic g) 
  thus ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Basic g, Stuck)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Spec r) 
  thus ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Spec r, Stuck)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Seq c1 c2)
  show ?case
  proof 
    assume "Γ⊢ (Seq c1 c2, Stuck) → …(∞)"
    from split_inf_Seq [OF this] Seq.hyps
    show False
      by (auto dest: steps_Stuck_prop)
  qed
next
  case (Cond b c1 c2) 
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Cond b c1 c2, Stuck)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (While b c) 
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (While b c, Stuck)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Call p) 
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Call p, Stuck)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (DynCom d) 
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (DynCom d, Stuck)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Guard m g c)
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Guard m g c, Stuck)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case Throw
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Throw, Stuck)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Catch c1 c2)
  show ?case
  proof 
    assume "Γ⊢ (Catch c1 c2, Stuck) → …(∞)"
    from split_inf_Catch [OF this] Catch.hyps
    show False
      by (auto dest: steps_Stuck_prop)
  qed  
next 
  case (Await b c)
  show ?case 
   proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Await b c, Stuck)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
qed

lemma not_inf_Fault: "¬ Γ⊢(c,Fault x) → …(∞)"
proof (induct c)
  case Skip
  show ?case
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Skip, Fault x)" 
    from f_step [of 0] f_0
    show False
      by (auto elim: Skip_no_step)
  qed
next
  case (Basic g) 
  thus ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Basic g, Fault x)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Spec r) 
  thus ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Spec r, Fault x)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Seq c1 c2)
  show ?case
  proof 
    assume "Γ⊢ (Seq c1 c2, Fault x) → …(∞)"
    from split_inf_Seq [OF this] Seq.hyps
    show False
      by (auto dest: steps_Fault_prop)
  qed
next
  case (Cond b c1 c2) 
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Cond b c1 c2, Fault x)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (While b c) 
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (While b c, Fault x)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Call p) 
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Call p, Fault x)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (DynCom d) 
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (DynCom d, Fault x)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Guard m g c)
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Guard m g c, Fault x)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case Throw
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Throw, Fault x)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Catch c1 c2)
  show ?case
  proof 
    assume "Γ⊢ (Catch c1 c2, Fault x) → …(∞)"
    from split_inf_Catch [OF this] Catch.hyps
    show False
      by (auto dest: steps_Fault_prop)
  qed  
next
  case (Await b c) 
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Await b c, Fault x)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
qed

lemma not_inf_Abrupt: "¬ Γ⊢(c,Abrupt s) → …(∞)"
proof (induct c)
  case Skip
  show ?case
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Skip, Abrupt s)" 
    from f_step [of 0] f_0
    show False
      by (auto elim: Skip_no_step)
  qed
next
  case (Basic g) 
  thus ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Basic g, Abrupt s)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Spec r) 
  thus ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Spec r, Abrupt s)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Seq c1 c2)
  show ?case
  proof 
    assume "Γ⊢ (Seq c1 c2, Abrupt s) → …(∞)"
    from split_inf_Seq [OF this] Seq.hyps
    show False
      by (auto dest: steps_Abrupt_prop)
  qed
next
  case (Cond b c1 c2) 
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Cond b c1 c2, Abrupt s)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (While b c) 
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (While b c, Abrupt s)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Call p) 
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Call p, Abrupt s)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (DynCom d) 
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (DynCom d, Abrupt s)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Guard m g c)
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Guard m g c, Abrupt s)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case Throw
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Throw, Abrupt s)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Catch c1 c2)
  show ?case
  proof 
    assume "Γ⊢ (Catch c1 c2, Abrupt s) → …(∞)"
    from split_inf_Catch [OF this] Catch.hyps
    show False
      by (auto dest: steps_Abrupt_prop)
  qed  
case (Await b c) 
  show ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Await b c, Abrupt s)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
qed


theorem terminates_impl_no_infinite_computation:
  assumes termi: "Γ⊢c ↓ s"
  shows "¬ Γ⊢(c,s) → …(∞)"
using termi
proof (induct)
  case (Skip s) thus ?case
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Skip, Normal s)" 
    from f_step [of 0] f_0
    show False
      by (auto elim: Skip_no_step)
  qed
next
  case (Basic g s) 
  thus ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Basic g, Normal s)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Spec r s) 
  thus ?case 
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Spec r, Normal s)" 
    from f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Guard s g c m)
  have g: "s ∈ g" by fact
  have hyp: "¬ Γ⊢ (c, Normal s) → …(∞)" by fact
  show ?case
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Guard m g c, Normal s)" 
    from f_step [of 0] f_0  g
    have "f 1 = (c,Normal s)"
      by (fastforce elim: step_elim_cases)
    with f_step
    have "Γ⊢ (c, Normal s) → …(∞)"
      apply (simp add: inf_def)
      apply (rule_tac x="λi. f (Suc i)" in exI)
      by simp
    with hyp show False ..
  qed
next
  case (GuardFault s g m c)
  have g: "s ∉ g" by fact
  show ?case
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Guard m g c, Normal s)" 
    from g f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Fault c m) 
  thus ?case
    by (rule not_inf_Fault)
next
  case (Seq c1 s c2)
  show ?case
  proof 
    assume "Γ⊢ (Seq c1 c2, Normal s) → …(∞)"
    from split_inf_Seq [OF this] Seq.hyps
    show False
      by (auto intro: steps_Skip_impl_exec)
  qed
next
  case (CondTrue s b c1 c2)
  have b: "s ∈ b" by fact
  have hyp_c1: "¬ Γ⊢ (c1, Normal s) → …(∞)" by fact
  show ?case
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Cond b c1 c2, Normal s)" 
    from b f_step [of 0] f_0
    have "f 1 = (c1,Normal s)"
      by (auto elim: step_Normal_elim_cases)
    with f_step
    have "Γ⊢ (c1, Normal s) → …(∞)"
      apply (simp add: inf_def)
      apply (rule_tac x="λi. f (Suc i)" in exI)
      by simp
    with hyp_c1 show False by simp
  qed    
next
  case (CondFalse s b c2 c1)
  have b: "s ∉ b" by fact
  have hyp_c2: "¬ Γ⊢ (c2, Normal s) → …(∞)" by fact
  show ?case
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Cond b c1 c2, Normal s)" 
    from b f_step [of 0] f_0
    have "f 1 = (c2,Normal s)"
      by (auto elim: step_Normal_elim_cases)
    with f_step
    have "Γ⊢ (c2, Normal s) → …(∞)"
      apply (simp add: inf_def)
      apply (rule_tac x="λi. f (Suc i)" in exI)
      by simp
    with hyp_c2 show False by simp
  qed
next    
  case (WhileTrue s b c)
  have b: "s ∈ b" by fact
  have hyp_c: "¬ Γ⊢ (c, Normal s) → …(∞)" by fact
  have hyp_w: "∀s'. Γ⊢ ⟨c,Normal s⟩ ⇒ s' ⟶ 
                     Γ⊢While b c ↓ s' ∧ ¬ Γ⊢ (While b c, s') → …(∞)" by fact
  have not_inf_Seq: "¬ Γ⊢ (Seq c (While b c), Normal s) → …(∞)"
  proof 
    assume "Γ⊢ (Seq c (While b c), Normal s) → …(∞)"
    from split_inf_Seq [OF this] hyp_c hyp_w show False
      by (auto intro: steps_Skip_impl_exec)
  qed
  show ?case
  proof 
    assume "Γ⊢ (While b c, Normal s) → …(∞)"
    then obtain f where
      f_step: "⋀i. Γ⊢f i → f (Suc i)" and
      f_0: "f 0 = (While b c, Normal s)" 
      by (auto simp add: inf_def)
    from f_step [of 0] f_0 b
    have "f 1 = (Seq c (While b c),Normal s)"
      by (auto elim: step_Normal_elim_cases)
    with f_step
    have "Γ⊢ (Seq c (While b c), Normal s) → …(∞)"
      apply (simp add: inf_def)
      apply (rule_tac x="λi. f (Suc i)" in exI)
      by simp
    with not_inf_Seq show False by simp
  qed
next
  case (WhileFalse s b c)
  have b: "s ∉ b" by fact
  show ?case
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (While b c, Normal s)" 
    from b f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Call p bdy s)
  have bdy: "Γ p = Some bdy" by fact
  have hyp: "¬ Γ⊢ (bdy, Normal s) → …(∞)" by fact
  show ?case
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Call p, Normal s)" 
    from bdy f_step [of 0] f_0
    have "f 1 = (bdy,Normal s)"
      by (auto elim: step_Normal_elim_cases)
    with f_step
    have "Γ⊢ (bdy, Normal s) → …(∞)"
      apply (simp add: inf_def)
      apply (rule_tac x="λi. f (Suc i)" in exI)
      by simp
    with hyp show False by simp
  qed    
next
  case (CallUndefined p s)
  have no_bdy: "Γ p = None" by fact
  show ?case
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Call p, Normal s)" 
    from no_bdy f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
next
  case (Stuck c)
  show ?case
    by (rule not_inf_Stuck)
next
  case (DynCom c s)
  have hyp: "¬ Γ⊢ (c s, Normal s) → …(∞)" by fact
  show ?case
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (DynCom c, Normal s)" 
    from f_step [of 0] f_0
    have "f (Suc 0) = (c s, Normal s)"
      by (auto elim: step_elim_cases)
    with f_step have "Γ⊢ (c s, Normal s) → …(∞)"
      apply (simp add: inf_def)
      apply (rule_tac x="λi. f (Suc i)" in exI)
      by simp
    with hyp
    show False by simp
  qed
next (Γ,c) ∝ clist
  case (Throw s) thus ?case
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Throw, Normal s)" 
    from f_step [of 0] f_0
    show False
      by (auto elim: step_elim_cases)
  qed  
next
  case (Abrupt c)
  show ?case
    by (rule not_inf_Abrupt)
next
  case (Catch c1 s c2)
  show ?case
  proof 
    assume "Γ⊢ (Catch c1 c2, Normal s) → …(∞)"
    from split_inf_Catch [OF this] Catch.hyps
    show False
      by (auto intro: steps_Throw_impl_exec)
  qed
next
  case (AwaitTrue s b c)
  have b: "s ∈ b" by fact
  show ?case
  proof (rule not_infI)
    fix f
    assume f_step: "⋀i. Γ⊢f i → f (Suc i)"
    assume f_0: "f 0 = (Await b c, Normal s)" 
    from b f_step [of 0] f_0 f_step [of 1]
    show False
      by (fastforce elim: Skip_no_step step_elim_cases)
  qed
qed


definition
 termi_call_steps :: "('s,'p,'f,'e) body ⇒ (('s × 'p) × ('s × 'p))set"
where
"termi_call_steps Γ =
 {((t,q),(s,p)). Γ⊢Call p↓Normal s ∧ 
       (∃c. Γ⊢(Call p,Normal s) →+ (c,Normal t) ∧ redex c = Call q)}"


primrec subst_redex:: "('s,'p,'f,'e)com ⇒ ('s,'p,'f,'e)com ⇒ ('s,'p,'f,'e)com"
where
"subst_redex Skip c = c" |
"subst_redex (Basic f) c = c" |
"subst_redex (Spec r) c = c" |
"subst_redex (Seq c1 c2) c  = Seq (subst_redex c1 c) c2" |
"subst_redex (Cond b c1 c2) c = c" |
"subst_redex (While b c') c = c" |
"subst_redex (Await b c') c = c" |
"subst_redex (Call p) c = c" |
"subst_redex (DynCom d) c = c" |
"subst_redex (Guard f b c') c = c" |
"subst_redex (Throw) c = c" |
"subst_redex (Catch c1 c2) c = Catch (subst_redex c1 c) c2"

lemma subst_redex_redex:
  "subst_redex c (redex c) = c"
  by (induct c) auto

lemma redex_subst_redex: "redex (subst_redex c r) = redex r"
  by (induct c) auto
  
lemma step_redex':
  shows "Γ⊢(redex c,s) → (r',s') ⟹ Γ⊢(c,s) → (subst_redex c r',s')"
by (induct c) (auto intro: step.Seq step.Catch)


lemma step_redex:
  shows "Γ⊢(r,s) → (r',s') ⟹ Γ⊢(subst_redex c r,s) → (subst_redex c r',s')"
by (induct c) (auto intro: step.Seq step.Catch)

lemma steps_redex:
  assumes steps: "Γ⊢ (r, s) →* (r', s')"
  shows "⋀c. Γ⊢(subst_redex c r,s) →* (subst_redex c r',s')"
using steps
proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans])
  case Refl 
  show "Γ⊢ (subst_redex c r', s') →* (subst_redex c r', s')"
    by simp
next
  case (Trans r s r'' s'')
  have "Γ⊢ (r, s) → (r'', s'')" by fact
  from step_redex [OF this]
  have "Γ⊢ (subst_redex c r, s) → (subst_redex c r'', s'')".
  also  
  have "Γ⊢ (subst_redex c r'', s'') →* (subst_redex c r', s')" by fact
  finally show ?case .
qed

ML {*
  ML_Thms.bind_thm ("trancl_induct2", Split_Rule.split_rule @{context}
    (Rule_Insts.read_instantiate @{context}
      [((("a", 0), Position.none), "(aa, ab)"), ((("b", 0), Position.none), "(ba, bb)")] []
      @{thm trancl_induct}));
*}

lemma steps_redex':
  assumes steps: "Γ⊢ (r, s) →+ (r', s')"
  shows "⋀c. Γ⊢(subst_redex c r,s) →+ (subst_redex c r',s')"
using steps
proof (induct rule: tranclp_induct2 [consumes 1,case_names Step Trans])
  case (Step r' s')
  have "Γ⊢ (r, s) → (r', s')" by fact
  then have "Γ⊢ (subst_redex c r, s) → (subst_redex c r', s')"
    by (rule step_redex)
  then show "Γ⊢ (subst_redex c r, s) →+ (subst_redex c r', s')"..
next
  case (Trans r' s' r'' s'')
  have "Γ⊢ (subst_redex c r, s) →+ (subst_redex c r', s')" by fact
  also
  have "Γ⊢ (r', s') → (r'', s'')" by fact
  hence "Γ⊢ (subst_redex c r', s') → (subst_redex c r'', s'')"
    by (rule step_redex)
  finally show "Γ⊢ (subst_redex c r, s) →+ (subst_redex c r'', s'')" .
qed

primrec seq:: "(nat ⇒ ('s,'p,'f,'e)com) ⇒ 'p ⇒ nat ⇒ ('s,'p,'f,'e)com"
where
"seq c p 0 = Call p" |
"seq c p (Suc i) = subst_redex (seq c p i) (c i)"


lemma renumber':
  assumes f: "∀i. (a,f i) ∈ r* ∧ (f i,f(Suc i)) ∈ r" 
  assumes a_b: "(a,b) ∈ r*" 
  shows "b = f 0 ⟹ (∃f. f 0 = a ∧ (∀i. (f i, f(Suc i)) ∈ r))"
using a_b
proof (induct rule: converse_rtrancl_induct [consumes 1])
  assume "b = f 0"
  with f show "∃f. f 0 = b ∧ (∀i. (f i, f (Suc i)) ∈ r)"
    by blast
next
  fix a z
  assume a_z: "(a, z) ∈ r" and "(z, b) ∈ r*" 
  assume "b = f 0 ⟹ ∃f. f 0 = z ∧ (∀i. (f i, f (Suc i)) ∈ r)"
         "b = f 0"
  then obtain f where f0: "f 0 = z" and seq: "∀i. (f i, f (Suc i)) ∈ r"
    by iprover
  {
    fix i have "((λi. case i of 0 ⇒ a | Suc i ⇒ f i) i, f i) ∈ r"
      using seq a_z f0
      by (cases i) auto
  }
  then
  show "∃f. f 0 = a ∧ (∀i. (f i, f (Suc i)) ∈ r)"
    by - (rule exI [where x="λi. case i of 0 ⇒ a | Suc i ⇒ f i"],simp)
qed

lemma renumber:
 "∀i. (a,f i) ∈ r* ∧ (f i,f(Suc i)) ∈ r 
 ⟹ ∃f. f 0 = a ∧ (∀i. (f i, f(Suc i)) ∈ r)"
  by (blast dest:renumber')

lemma lem:
  "∀y. r++ a y ⟶ P a ⟶ P y 
   ⟹ ((b,a) ∈ {(y,x). P x ∧ r x y}+) = ((b,a) ∈ {(y,x). P x ∧ r++ x y})"
apply(rule iffI)
 apply clarify
 apply(erule trancl_induct)
  apply blast
 apply(blast intro:tranclp_trans)
apply clarify
apply(erule tranclp_induct)
 apply blast
apply(blast intro:trancl_trans)
done

corollary terminates_impl_no_infinite_trans_computation:
 assumes terminates: "Γ⊢c↓s"
 shows "¬(∃f. f 0 = (c,s) ∧ (∀i. Γ⊢f i →+ f(Suc i)))"
proof -
  have "wf({(y,x). Γ⊢(c,s) →* x ∧ Γ⊢x → y}+)"
  proof (rule wf_trancl)
    show "wf {(y, x). Γ⊢(c,s) →* x ∧ Γ⊢x → y}"
    proof (simp only: wf_iff_no_infinite_down_chain,clarify,simp)
      fix f
      assume "∀i. Γ⊢(c,s) →* f i ∧ Γ⊢f i → f (Suc i)"
      hence "∃f. f (0::nat) = (c,s) ∧ (∀i. Γ⊢f i → f (Suc i))"
        by (rule renumber [to_pred])
      moreover from terminates_impl_no_infinite_computation [OF terminates]
      have "¬ (∃f. f (0::nat) = (c, s) ∧ (∀i. Γ⊢f i → f (Suc i)))"
        by (simp add: inf_def)
      ultimately show False
        by simp
    qed
  qed
  hence "¬ (∃f. ∀i. (f (Suc i), f i)
                 ∈ {(y, x). Γ⊢(c, s) →* x ∧ Γ⊢x → y}+)"
    by (simp add: wf_iff_no_infinite_down_chain)
  thus ?thesis
  proof (rule contrapos_nn)
    assume "∃f. f (0::nat) = (c, s) ∧ (∀i. Γ⊢f i →+ f (Suc i))"
    then obtain f where
      f0: "f 0 = (c, s)" and
      seq: "∀i. Γ⊢f i →+ f (Suc i)"
      by iprover
    show 
      "∃f. ∀i. (f (Suc i), f i) ∈ {(y, x). Γ⊢(c, s) →* x ∧ Γ⊢x → y}+"
    proof (rule exI [where x=f],rule allI)
      fix i
      show "(f (Suc i), f i) ∈ {(y, x). Γ⊢(c, s) →* x ∧ Γ⊢x → y}+"
      proof -   
        {
          fix i have "Γ⊢(c,s) →* f i"
          proof (induct i)
            case 0 show "Γ⊢(c, s) →* f 0"
              by (simp add: f0)
          next
            case (Suc n)
            have "Γ⊢(c, s) →* f n"  by fact
            with seq show "Γ⊢(c, s) →* f (Suc n)"
              by (blast intro: tranclp_into_rtranclp rtranclp_trans)
          qed
        }
        hence "Γ⊢(c,s) →* f i"
          by iprover
        with seq have
          "(f (Suc i), f i) ∈ {(y, x). Γ⊢(c, s) →* x ∧ Γ⊢x →+ y}"
          by clarsimp
        moreover 
        have "∀y. Γ⊢f i →+ y⟶Γ⊢(c, s) →* f i⟶Γ⊢(c, s) →* y"
          by (blast intro: tranclp_into_rtranclp rtranclp_trans)
        ultimately 
        show ?thesis 
          by (subst lem )
      qed
    qed
  qed
qed

theorem wf_termi_call_steps: "wf (termi_call_steps Γ)"
proof (simp only: termi_call_steps_def wf_iff_no_infinite_down_chain,
       clarify,simp)
  fix f
  assume inf: "∀i. (λ(t, q) (s, p).
                Γ⊢Call p ↓ Normal s ∧
                (∃c. Γ⊢ (Call p, Normal s) →+ (c, Normal t) ∧ redex c = Call q))
             (f (Suc i)) (f i)"
  def s≡"λi::nat. fst (f i)" 
  def p≡"λi::nat. snd (f i)::'b"
  from inf
  have inf': "∀i. Γ⊢Call (p i) ↓ Normal (s i) ∧
               (∃c. Γ⊢ (Call (p i), Normal (s i)) →+ (c, Normal (s (i+1))) ∧ 
                    redex c = Call (p (i+1)))"
    apply -
    apply (rule allI)
    apply (erule_tac x=i in allE)
    apply (auto simp add: s_def p_def)
    done
  show False
  proof -
    from inf'
    have "∃c. ∀i. Γ⊢Call (p i) ↓ Normal (s i) ∧
               Γ⊢ (Call (p i), Normal (s i)) →+ (c i, Normal (s (i+1))) ∧ 
                    redex (c i) = Call (p (i+1))"
      apply -
      apply (rule choice)
      by blast
    then obtain c where
      termi_c: "∀i. Γ⊢Call (p i) ↓ Normal (s i)" and
      steps_c: "∀i. Γ⊢ (Call (p i), Normal (s i)) →+ (c i, Normal (s (i+1)))" and
      red_c:   "∀i. redex (c i) = Call (p (i+1))"
      by auto
    def g≡"λi. (seq c (p 0) i,Normal (s i)::('a,'c) xstate)"
    from red_c [rule_format, of 0]
    have "g 0 = (Call (p 0), Normal (s 0))"
      by (simp add: g_def)
    moreover
    {
      fix i
      have "redex (seq c (p 0) i) = Call (p i)"
        by (induct i) (auto simp add: redex_subst_redex red_c)
      from this [symmetric]
      have "subst_redex (seq c (p 0) i) (Call (p i)) = (seq c (p 0) i)"
        by (simp add: subst_redex_redex)
    } note subst_redex_seq = this
    have "∀i. Γ⊢ (g i) →+ (g (i+1))"
    proof 
      fix i
      from steps_c [rule_format, of i]
      have "Γ⊢ (Call (p i), Normal (s i)) →+ (c i, Normal (s (i + 1)))".
      from steps_redex' [OF this, of "(seq c (p 0) i)"]
      have "Γ⊢ (subst_redex (seq c (p 0) i) (Call (p i)), Normal (s i)) →+
                (subst_redex (seq c (p 0) i) (c i), Normal (s (i + 1)))" .
      hence "Γ⊢ (seq c (p 0) i, Normal (s i)) →+ 
                 (seq c (p 0) (i+1), Normal (s (i + 1)))"
        by (simp add: subst_redex_seq)
      thus "Γ⊢ (g i) →+ (g (i+1))"
        by (simp add: g_def)
    qed
    moreover
    from terminates_impl_no_infinite_trans_computation [OF termi_c [rule_format, of 0]]
    have "¬ (∃f. f 0 = (Call (p 0), Normal (s 0)) ∧ (∀i. Γ⊢ f i →+ f (Suc i)))" .
    ultimately show False
      by auto
  qed
qed


lemma no_infinite_computation_implies_wf: 
  assumes not_inf: "¬ Γ⊢ (c, s) → …(∞)"
  shows "wf {(c2,c1). Γ ⊢ (c,s) →* c1 ∧ Γ ⊢ c1 → c2}"
proof (simp only: wf_iff_no_infinite_down_chain,clarify, simp)
  fix f
  assume "∀i. Γ⊢(c, s) →* f i ∧ Γ⊢f i → f (Suc i)"
  hence "∃f. f 0 = (c, s) ∧ (∀i. Γ⊢f i → f (Suc i))"
    by (rule renumber [to_pred])
  moreover from not_inf
  have "¬ (∃f. f 0 = (c, s) ∧ (∀i. Γ⊢f i → f (Suc i)))"
    by (simp add: inf_def)
  ultimately show False
    by simp
qed

lemma not_final_Stuck_step: "¬ final (c,Stuck) ⟹ ∃c' s'. Γ⊢ (c, Stuck) → (c',s')"
by (induct c) (fastforce intro: step.intros simp add: final_def)+

lemma not_final_Abrupt_step: 
  "¬ final (c,Abrupt s) ⟹ ∃c' s'. Γ⊢ (c, Abrupt s) → (c',s')"
by (induct c) (fastforce intro: step.intros simp add: final_def)+

lemma not_final_Fault_step: 
  "¬ final (c,Fault f) ⟹ ∃c' s'. Γ⊢ (c, Fault f) → (c',s')"
by (induct c) (fastforce intro: step.intros simp add: final_def)+

lemma not_final_Normal_step: 
  "¬ final (c,Normal s) ∧ ((∃b c1. redex c = Await b c1) ⟶ Γ⊢c ↓ Normal s) ⟹ ∃c' s'. Γ⊢ (c, Normal s) → (c',s')"
proof (induct c) 
  case Skip thus ?case by (simp add: final_def)
next
  case Basic thus ?case by (meson step.Basic)
next
  case (Spec r)
  thus ?case by (meson step.Spec step.SpecStuck)    
next
  case (Seq c1 c2)
  thus ?case by (metis SeqSkip SeqThrow final_def fst_conv redex.simps(4) step.Seq terminates_Normal_elim_cases(5))
next
  case (Cond b c1 c2)
  show ?case
    by (cases "s ∈ b") (fastforce intro: step.intros)+
next
  case (While b c)
  show ?case
    by (cases "s ∈ b") (fastforce intro: step.intros)+
next
  case (Call p)
  show ?case
  by (cases "Γ p") (fastforce intro: step.intros)+
next
  case DynCom thus ?case by (fastforce intro: step.intros)
next
  case (Guard f g c)
  show ?case
    by (cases "s ∈ g") (fastforce intro: step.intros)+
next
  case Throw
  thus ?case by (fastforce intro: step.intros simp add: final_def)
next
  case (Catch c1 c2)
  thus ?case
    by (cases "final (c1,Normal s)") (fastforce intro: step.intros elim: terminates_Normal_elim_cases simp add: final_def)+
next
  case (Await b c) 
  then obtain ba c1 where x:"(∃ba c1. redex (Await b c) = Await ba c1)" by simp
  then have "Γ⊢Await b c ↓ Normal s" using x Await.prems by blast 
  also have "s ∈ b" by (meson `Γ⊢Await b c ↓ Normal s` terminates_Normal_elim_cases(12))
  moreover have "∃t. Γ⊢⟨c, Normal s⟩ ⇒ t" by (meson calculation terminates_Normal_elim_cases(12) terminates_implies_exec) 
  ultimately show ?case using AwaitAbrupt  step.Await by fastforce
qed

lemma final_termi:
"final (c,s) ⟹ Γ⊢c↓s"
  by (cases s) (auto simp add: final_def terminates.intros)

lemma split_computation: 
assumes steps: "Γ⊢ (c, s) →* (cf, sf)"
assumes not_final: "¬ final (c,s)"
assumes final: "final (cf,sf)"
shows "∃c' s'. Γ⊢ (c, s) → (c',s') ∧ Γ⊢ (c', s') →* (cf, sf)"
using steps not_final final
proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans])
  case Refl thus ?case by simp
next
  case (Trans c s c' s')
  thus ?case by auto
qed


lemma wf_implies_termi_reach_step_case:
assumes hyp: "⋀c' s'. Γ⊢ (c, Normal s) → (c', s')⟹ Γ⊢c' ↓ s'" and
        hyp1:"(⋀b c1. (redex c = Await b c1) ⟹ Γ⊢c1 ↓ Normal s ∧ s∈b)"
shows "Γ⊢c ↓ Normal s"
using hyp hyp1
proof (induct c)
  case Skip show ?case by (fastforce intro: terminates.intros)
next
  case Basic show ?case by (fastforce intro: terminates.intros)
next
  case (Spec r)
  show ?case
    by (cases "∃t. (s,t)∈r") (fastforce intro: terminates.intros)+
next
  case (Seq c1 c2)
  have hyp: "⋀c' s'. Γ⊢ (Seq c1 c2, Normal s) → (c', s') ⟹ Γ⊢c' ↓ s'"
using Seq.prems by blast 
   have hyp': "(⋀b c1. (redex (Seq c1 c2)  = Await b c1) ⟹ Γ⊢c1 ↓ Normal s ∧ s∈b)" by fact 
  show ?case 
  proof (rule terminates.Seq)
    {
      fix c' s'
      assume step_c1: "Γ⊢ (c1, Normal s) → (c', s')"      
      assume red: "(⋀b c1. (redex c1  = Await b c1) ⟹ Γ⊢c1 ↓ Normal s ∧ s∈b)"
      have "Γ⊢c' ↓ s'"
      proof -
        from step_c1
        have "Γ⊢ (Seq c1 c2, Normal s) → (Seq c' c2, s')"
          by (simp add: step.Seq)          
        from hyp [OF this]
        have "Γ⊢Seq c' c2 ↓ s'".
        thus "Γ⊢c'↓ s'"
          by cases auto
      qed
    } 
    from Seq.hyps (1) [OF this]
    show "Γ⊢c1 ↓ Normal s" using terminates_Skip' by (simp add: hyp') 
  next
    show "∀s'. Γ⊢ ⟨c1,Normal s⟩ ⇒ s' ⟶ Γ⊢c2 ↓ s'"
    proof (intro allI impI)
      fix s'
      assume exec_c1: "Γ⊢ ⟨c1,Normal s⟩ ⇒ s'"
      show "Γ⊢c2 ↓ s'"
      proof (cases "final (c1,Normal s)")
        case True
        hence "c1=Skip ∨ c1=Throw"
          by (simp add: final_def)
        thus ?thesis
        proof
          assume Skip: "c1=Skip"
          have s1:"Γ⊢(Seq Skip c2,Normal s) → (c2,Normal s) ∧¬ (∃b c1. Seq c1 c2 = Await b c1)"
            by (simp add: step.SeqSkip)          
          from hyp 
          have "Γ⊢c2 ↓ Normal s" using local.Skip s1 by blast 
          moreover from exec_c1 Skip
          have "s'=Normal s"
            by (auto elim: exec_Normal_elim_cases)
          ultimately show ?thesis by simp
        next
          assume Throw: "c1=Throw"
          with exec_c1 have "s'=Abrupt s"
            by (auto elim: exec_Normal_elim_cases)
          thus ?thesis
            by auto
        qed
      next
        case False        
        from exec_impl_steps [OF exec_c1]
        obtain cf t where 
          steps_c1: "Γ⊢ (c1, Normal s) →* (cf, t)" and
          fin:"(case s' of
                 Abrupt x ⇒ cf = Throw ∧ t = Normal x
                | _ ⇒ cf = Skip ∧ t = s')"
          by (fastforce split: xstate.splits)
        with fin have final: "final (cf,t)"
          by (cases s') (auto simp add: final_def)
        from split_computation [OF steps_c1 False this]
        obtain c'' s'' where
          first: "Γ⊢ (c1, Normal s) → (c'', s'')" and
          rest: "Γ⊢ (c'', s'') →* (cf, t)" 
          by blast
        from step.Seq [OF first]
        have "Γ⊢ (Seq c1 c2, Normal s) → (Seq c'' c2, s'')" by auto
        from hyp [OF this]
        have termi_s'': "Γ⊢Seq c'' c2 ↓ s''".
        show ?thesis
        proof (cases s'')
          case (Normal x)
          from termi_s'' [simplified Normal]
          have termi_c2: "∀t. Γ⊢ ⟨c'',Normal x⟩ ⇒ t ⟶ Γ⊢c2 ↓ t"
            by cases
          show ?thesis
          proof (cases "∃x'. s'=Abrupt x'")
            case False
            with fin obtain "cf=Skip" "t=s'"
              by (cases s') auto
            from steps_Skip_impl_exec [OF rest [simplified this]] Normal
            have "Γ⊢ ⟨c'',Normal x⟩ ⇒ s'"
              by simp
            from termi_c2 [rule_format, OF this]
            show "Γ⊢c2 ↓ s'" .
          next
            case True
            with fin obtain x' where s': "s'=Abrupt x'" and "cf=Throw" "t=Normal x'"
              by auto
            from steps_Throw_impl_exec [OF rest [simplified this]] Normal 
            have "Γ⊢ ⟨c'',Normal x⟩ ⇒ Abrupt x'"
              by simp
            from termi_c2 [rule_format, OF this] s'
            show "Γ⊢c2 ↓ s'" by simp
           qed
        next
          case (Abrupt x)
          from steps_Abrupt_prop [OF rest this]
          have "t=Abrupt x" by simp
          with fin have "s'=Abrupt x"
            by (cases s') auto
          thus "Γ⊢c2 ↓ s'" 
            by auto
        next
          case (Fault f)
          from steps_Fault_prop [OF rest this]
          have "t=Fault f" by simp
          with fin have "s'=Fault f"
            by (cases s') auto
          thus "Γ⊢c2 ↓ s'" 
            by auto
        next
          case Stuck
          from steps_Stuck_prop [OF rest this]
          have "t=Stuck" by simp
          with fin have "s'=Stuck"
            by (cases s') auto
          thus "Γ⊢c2 ↓ s'" 
            by auto
        qed
      qed
    qed
  qed
next
  case (Cond b c1 c2)
  have hyp: "⋀c' s'. Γ⊢ (Cond b c1 c2, Normal s) → (c', s') ⟹ Γ⊢c' ↓ s'"
by (simp add: Cond.prems(1)) 
  show ?case
  proof (cases "s∈b") 
    case True
    then have "Γ⊢ (Cond b c1 c2, Normal s) → (c1, Normal s) "
     by (simp add: step.CondTrue)     
    from hyp [OF this] have "Γ⊢c1 ↓ Normal s".
    with True show ?thesis
      by (auto intro: terminates.intros)
  next
    case False
    then have "Γ⊢ (Cond b c1 c2, Normal s) → (c2, Normal s)"
      by (simp add: step.CondFalse)
    from hyp [OF this] have "Γ⊢c2 ↓ Normal s".
    with False show ?thesis
      by (auto intro: terminates.intros)
  qed
next
  case (While b c)
  have hyp: "⋀c' s'. Γ⊢ (While b c, Normal s) → (c', s')   ⟹ Γ⊢c' ↓ s'" by fact
  show ?case
  proof (cases "s∈b") 
    case True
    then have "Γ⊢ (While b c, Normal s) → (Seq c (While b c), Normal s)"
      by (simp add: step.WhileTrue)
    from hyp [OF this] have "Γ⊢(Seq c (While b c)) ↓ Normal s".
    with True show ?thesis
      by (auto elim: terminates_Normal_elim_cases intro: terminates.intros)
  next
    case False
    thus ?thesis
      by (auto intro: terminates.intros)
  qed
next
  case (Call p)
  have hyp: "⋀c' s'. Γ⊢ (Call p, Normal s) → (c', s')  ⟹ Γ⊢c' ↓ s'" by (simp add: Call.prems)
  show ?case
  proof (cases "Γ p")
    case None
    thus ?thesis
      by (auto intro: terminates.intros)
  next
    case (Some bdy)
    then have "Γ⊢ (Call p, Normal s) → (bdy, Normal s)"
      by (simp add: step.Call)
    from hyp [OF this] have "Γ⊢bdy ↓ Normal s".
    with Some show ?thesis
      by (auto intro: terminates.intros)
  qed
next
  case (DynCom c)
  have hyp: "⋀c' s'. Γ⊢ (DynCom c, Normal s) → (c', s')  ⟹ Γ⊢c' ↓ s'" by fact
  have "Γ⊢ (DynCom c, Normal s) → (c s, Normal s)"
    by (simp add: step.DynCom)
  from hyp [OF this] have "Γ⊢c s ↓ Normal s".
  then show ?case
    by (auto intro: terminates.intros)
next
  case (Guard f g c)
  have hyp: "⋀c' s'. Γ⊢ (Guard f g c, Normal s) → (c', s')   ⟹ Γ⊢c' ↓ s'" by fact
  show ?case
  proof (cases "s∈g")
    case True
    then have "Γ⊢ (Guard f g c, Normal s) → (c, Normal s) "
      by (simp add: step.Guard) thm step.Guard
    from hyp [OF this] have "Γ⊢c↓ Normal s".
    with True show ?thesis
      by (auto intro: terminates.intros)
  next
    case False
    thus ?thesis
      by (auto intro: terminates.intros)
  qed
next 
  case Throw show ?case by (auto intro: terminates.intros)
next
  case (Catch c1 c2)
  have hyp: "⋀c' s'. Γ⊢ (Catch c1 c2, Normal s) → (c', s')   ⟹ Γ⊢c' ↓ s'" by fact
  have hyp': "(⋀b c1. (redex (Catch c1 c2) = Await b c1) ⟹ Γ⊢c1 ↓ Normal s ∧ s∈b)" by fact
  show ?case
  proof (rule terminates.Catch)
    {
      fix c' s'
      assume step_c1: "Γ⊢ (c1, Normal s) → (c', s')"
      assume red: "(⋀b c1. (redex c1  = Await b c1) ⟹ Γ⊢c1 ↓ Normal s ∧ s∈b)"
      have "Γ⊢c' ↓ s'"
      proof -
        from step_c1
        have "Γ⊢ (Catch c1 c2, Normal s) → (Catch c' c2, s') "
          by (simp add: step.Catch)
        from hyp [OF this]
        have "Γ⊢Catch c' c2 ↓ s'".
        thus "Γ⊢c'↓ s'"
          by cases auto
      qed
    } 
    from Catch.hyps (1) [OF this]
    show "Γ⊢c1 ↓ Normal s" using terminates_Skip' by (simp add: hyp')
  next  
    show "∀s'. Γ⊢ ⟨c1,Normal s⟩ ⇒ Abrupt s' ⟶ Γ⊢c2 ↓ Normal s'"
    proof (intro allI impI)
      fix s'
      assume exec_c1: "Γ⊢ ⟨c1,Normal s⟩ ⇒ Abrupt s'"
      show "Γ⊢c2 ↓ Normal s'"
      proof (cases "final (c1,Normal s)")
        case True
        with exec_c1
        have Throw: "c1=Throw"
          by (auto simp add: final_def elim: exec_Normal_elim_cases)
        have s1: "Γ⊢(Catch Throw c2,Normal s) → (c2,Normal s)"
          by (simp add: step.CatchThrow)
        from hyp 
        have "Γ⊢c2 ↓ Normal s" using local.Throw s1 by blast 
        moreover from exec_c1 Throw
        have "s'=s"
          by (auto elim: exec_Normal_elim_cases)
        ultimately show ?thesis by simp
      next
        case False
        from exec_impl_steps [OF exec_c1]
        obtain cf t where 
          steps_c1: "Γ⊢ (c1, Normal s) →* (Throw, Normal s')" 
          by (fastforce split: xstate.splits)
        from split_computation [OF steps_c1 False]
        obtain c'' s'' where
          first: "Γ⊢ (c1, Normal s) → (c'', s'')" and
          rest: "Γ⊢ (c'', s'') →* (Throw, Normal s')" 
          by (auto simp add: final_def)
        from step.Catch [OF first]
        have "Γ⊢ (Catch c1 c2, Normal s) → (Catch c'' c2, s'')" by auto
        from hyp [OF this] 
        have "Γ⊢Catch c'' c2 ↓ s''".
        moreover
        from steps_Throw_impl_exec [OF rest]
        have "Γ⊢ ⟨c'',s''⟩ ⇒ Abrupt s'".
        moreover
        from rest obtain x where "s''=Normal x"
          by (cases s'')
             (auto dest: steps_Fault_prop steps_Abrupt_prop steps_Stuck_prop)
        ultimately show ?thesis
          by (fastforce elim: terminates_elim_cases)
      qed
    qed
  qed
next
  case (Await b c) 
  have hyp':"⋀ba c1. redex (Await b c) = Await ba c1 ⟹ Γ⊢c1 ↓ Normal s ∧ s ∈ ba" by fact 
  have hyp: "⋀c' s'. Γ⊢ (Await b c, Normal s) → (c', s') ⟹ Γ⊢c' ↓ s'" by (simp add: final_termi step_await_final1)
  have red:"redex (Await b c)  = Await b c" by auto  
  then have "⋀b1 c1. (redex (Await b c)  = Await b1 c1) ⟹ b1=b ∧ c1 = c" by simp
  then have " Γ⊢c ↓ Normal s ∧ s ∈ b" by (simp add: Await.prems(2))
  show ?case
  by (simp add: `Γ⊢c ↓ Normal s ∧ s ∈ b` terminates.AwaitTrue)      
qed


lemma wf_implies_termi_reach_step_case':
assumes hyp: "⋀c' s'. Γ⊢ (c, Normal s) → (c', s')⟹ Γ⊢c' ↓ s'" and
        hyp1:"~(∃b c1. redex c = Await b c1)"
shows "Γ⊢c ↓ Normal s"
using hyp hyp1
proof (induct c)
  case Skip show ?case by (fastforce intro: terminates.intros)
next
  case Basic show ?case by (fastforce intro: terminates.intros)
next
  case (Spec r)
  show ?case
    by (cases "∃t. (s,t)∈r") (fastforce intro: terminates.intros)+
next
  case (Seq c1 c2)
  have hyp: "⋀c' s'. Γ⊢ (Seq c1 c2, Normal s) → (c', s') ⟹ Γ⊢c' ↓ s'"
using Seq.prems by blast
  have hyp': "~(∃b c1. redex (Seq c1 c2) = Await b c1)" by fact
  show ?case 
  proof (rule terminates.Seq)
    {
      fix c' s'
      assume step_c1: "Γ⊢ (c1, Normal s) → (c', s')"
      assume red:  "~(∃b c1. redex c1 = Await b c1)"      
      have "Γ⊢c' ↓ s'"
      proof -
        from step_c1
        have "Γ⊢ (Seq c1 c2, Normal s) → (Seq c' c2, s')"
          by (simp add: step.Seq)          
        from hyp [OF this]
        have "Γ⊢Seq c' c2 ↓ s'".
        thus "Γ⊢c'↓ s'"
          by cases auto
      qed
    } 
    from Seq.hyps (1) [OF this]
    show "Γ⊢c1 ↓ Normal s"using hyp' by force 
  next
    show "∀s'. Γ⊢ ⟨c1,Normal s⟩ ⇒ s' ⟶ Γ⊢c2 ↓ s'"
    proof (intro allI impI)
      fix s'
      assume exec_c1: "Γ⊢ ⟨c1,Normal s⟩ ⇒ s'"
      show "Γ⊢c2 ↓ s'"
      proof (cases "final (c1,Normal s)")
        case True
        hence "c1=Skip ∨ c1=Throw"
          by (simp add: final_def)
        thus ?thesis
        proof
          assume Skip: "c1=Skip"
          have s1:"Γ⊢(Seq Skip c2,Normal s) → (c2,Normal s) ∧¬ (∃b c1. Seq c1 c2 = Await b c1)"
            by (simp add: step.SeqSkip)          
          from hyp 
          have "Γ⊢c2 ↓ Normal s" using local.Skip s1 by blast 
          moreover from exec_c1 Skip
          have "s'=Normal s"
            by (auto elim: exec_Normal_elim_cases)
          ultimately show ?thesis by simp
        next
          assume Throw: "c1=Throw"
          with exec_c1 have "s'=Abrupt s"
            by (auto elim: exec_Normal_elim_cases)
          thus ?thesis
            by auto
        qed
      next
        case False        
        from exec_impl_steps [OF exec_c1]
        obtain cf t where 
          steps_c1: "Γ⊢ (c1, Normal s) →* (cf, t)" and
          fin:"(case s' of
                 Abrupt x ⇒ cf = Throw ∧ t = Normal x
                | _ ⇒ cf = Skip ∧ t = s')"
          by (fastforce split: xstate.splits)
        with fin have final: "final (cf,t)"
          by (cases s') (auto simp add: final_def)
        from split_computation [OF steps_c1 False this]
        obtain c'' s'' where
          first: "Γ⊢ (c1, Normal s) → (c'', s'')" and
          rest: "Γ⊢ (c'', s'') →* (cf, t)" 
          by blast
        from step.Seq [OF first]
        have "Γ⊢ (Seq c1 c2, Normal s) → (Seq c'' c2, s'')" by auto
        from hyp [OF this]
        have termi_s'': "Γ⊢Seq c'' c2 ↓ s''".
        show ?thesis
        proof (cases s'')
          case (Normal x)
          from termi_s'' [simplified Normal]
          have termi_c2: "∀t. Γ⊢ ⟨c'',Normal x⟩ ⇒ t ⟶ Γ⊢c2 ↓ t"
            by cases
          show ?thesis
          proof (cases "∃x'. s'=Abrupt x'")
            case False
            with fin obtain "cf=Skip" "t=s'"
              by (cases s') auto
            from steps_Skip_impl_exec [OF rest [simplified this]] Normal
            have "Γ⊢ ⟨c'',Normal x⟩ ⇒ s'"
              by simp
            from termi_c2 [rule_format, OF this]
            show "Γ⊢c2 ↓ s'" .
          next
            case True
            with fin obtain x' where s': "s'=Abrupt x'" and "cf=Throw" "t=Normal x'"
              by auto
            from steps_Throw_impl_exec [OF rest [simplified this]] Normal 
            have "Γ⊢ ⟨c'',Normal x⟩ ⇒ Abrupt x'"
              by simp
            from termi_c2 [rule_format, OF this] s'
            show "Γ⊢c2 ↓ s'" by simp
           qed
        next
          case (Abrupt x)
          from steps_Abrupt_prop [OF rest this]
          have "t=Abrupt x" by simp
          with fin have "s'=Abrupt x"
            by (cases s') auto
          thus "Γ⊢c2 ↓ s'" 
            by auto
        next
          case (Fault f)
          from steps_Fault_prop [OF rest this]
          have "t=Fault f" by simp
          with fin have "s'=Fault f"
            by (cases s') auto
          thus "Γ⊢c2 ↓ s'" 
            by auto
        next
          case Stuck
          from steps_Stuck_prop [OF rest this]
          have "t=Stuck" by simp
          with fin have "s'=Stuck"
            by (cases s') auto
          thus "Γ⊢c2 ↓ s'" 
            by auto
        qed
      qed
    qed
  qed
next
  case (Cond b c1 c2)
  have hyp: "⋀c' s'. Γ⊢ (Cond b c1 c2, Normal s) → (c', s') ⟹ Γ⊢c' ↓ s'"
by (simp add: Cond.prems(1)) 
  show ?case
  proof (cases "s∈b") 
    case True
    then have "Γ⊢ (Cond b c1 c2, Normal s) → (c1, Normal s) "
     by (simp add: step.CondTrue)     
    from hyp [OF this] have "Γ⊢c1 ↓ Normal s".
    with True show ?thesis
      by (auto intro: terminates.intros)
  next
    case False
    then have "Γ⊢ (Cond b c1 c2, Normal s) → (c2, Normal s)"
      by (simp add: step.CondFalse)
    from hyp [OF this] have "Γ⊢c2 ↓ Normal s".
    with False show ?thesis
      by (auto intro: terminates.intros)
  qed
next
  case (While b c)
  have hyp: "⋀c' s'. Γ⊢ (While b c, Normal s) → (c', s')   ⟹ Γ⊢c' ↓ s'" by fact
  show ?case
  proof (cases "s∈b") 
    case True
    then have "Γ⊢ (While b c, Normal s) → (Seq c (While b c), Normal s)"
      by (simp add: step.WhileTrue)
    from hyp [OF this] have "Γ⊢(Seq c (While b c)) ↓ Normal s".
    with True show ?thesis
      by (auto elim: terminates_Normal_elim_cases intro: terminates.intros)
  next
    case False
    thus ?thesis
      by (auto intro: terminates.intros)
  qed
next
  case (Call p)
  have hyp: "⋀c' s'. Γ⊢ (Call p, Normal s) → (c', s')  ⟹ Γ⊢c' ↓ s'" by (simp add: Call.prems)
  show ?case
  proof (cases "Γ p")
    case None
    thus ?thesis
      by (auto intro: terminates.intros)
  next
    case (Some bdy)
    then have "Γ⊢ (Call p, Normal s) → (bdy, Normal s)"
      by (simp add: step.Call)
    from hyp [OF this] have "Γ⊢bdy ↓ Normal s".
    with Some show ?thesis
      by (auto intro: terminates.intros)
  qed
next
  case (DynCom c)
  have hyp: "⋀c' s'. Γ⊢ (DynCom c, Normal s) → (c', s')  ⟹ Γ⊢c' ↓ s'" by fact
  have "Γ⊢ (DynCom c, Normal s) → (c s, Normal s)"
    by (simp add: step.DynCom)
  from hyp [OF this] have "Γ⊢c s ↓ Normal s".
  then show ?case
    by (auto intro: terminates.intros)
next
  case (Guard f g c)
  have hyp: "⋀c' s'. Γ⊢ (Guard f g c, Normal s) → (c', s')   ⟹ Γ⊢c' ↓ s'" by fact
  show ?case
  proof (cases "s∈g")
    case True
    then have "Γ⊢ (Guard f g c, Normal s) → (c, Normal s) "
      by (simp add: step.Guard) thm step.Guard
    from hyp [OF this] have "Γ⊢c↓ Normal s".
    with True show ?thesis
      by (auto intro: terminates.intros)
  next
    case False
    thus ?thesis
      by (auto intro: terminates.intros)
  qed
next 
  case Throw show ?case by (auto intro: terminates.intros)
next
  case (Catch c1 c2)
  have hyp: "⋀c' s'. Γ⊢ (Catch c1 c2, Normal s) → (c', s')   ⟹ Γ⊢c' ↓ s'" by fact
  have hyp': "~(∃b c1. redex (Catch c1 c2) = Await b c1)" by fact
  show ?case
  proof (rule terminates.Catch)
    {
      fix c' s'
      assume step_c1: "Γ⊢ (c1, Normal s) → (c', s')"
      assume red: "~(∃b c1. redex c1  = Await b c1)"
      have "Γ⊢c' ↓ s'"
      proof -
        from step_c1
        have "Γ⊢ (Catch c1 c2, Normal s) → (Catch c' c2, s') "
          by (simp add: step.Catch)
        from hyp [OF this]
        have "Γ⊢Catch c' c2 ↓ s'".
        thus "Γ⊢c'↓ s'"
          by cases auto
      qed
    } 
    from Catch.hyps (1) [OF this]
    show "Γ⊢c1 ↓ Normal s" using hyp' by force
  next  
    show "∀s'. Γ⊢ ⟨c1,Normal s⟩ ⇒ Abrupt s' ⟶ Γ⊢c2 ↓ Normal s'"
    proof (intro allI impI)
      fix s'
      assume exec_c1: "Γ⊢ ⟨c1,Normal s⟩ ⇒ Abrupt s'"
      show "Γ⊢c2 ↓ Normal s'"
      proof (cases "final (c1,Normal s)")
        case True
        with exec_c1
        have Throw: "c1=Throw"
          by (auto simp add: final_def elim: exec_Normal_elim_cases)
        have s1: "Γ⊢(Catch Throw c2,Normal s) → (c2,Normal s)"
          by (simp add: step.CatchThrow)
        from hyp 
        have "Γ⊢c2 ↓ Normal s" using local.Throw s1 by blast 
        moreover from exec_c1 Throw
        have "s'=s"
          by (auto elim: exec_Normal_elim_cases)
        ultimately show ?thesis by simp
      next
        case False
        from exec_impl_steps [OF exec_c1]
        obtain cf t where 
          steps_c1: "Γ⊢ (c1, Normal s) →* (Throw, Normal s')" 
          by (fastforce split: xstate.splits)
        from split_computation [OF steps_c1 False]
        obtain c'' s'' where
          first: "Γ⊢ (c1, Normal s) → (c'', s'')" and
          rest: "Γ⊢ (c'', s'') →* (Throw, Normal s')" 
          by (auto simp add: final_def)
        from step.Catch [OF first]
        have "Γ⊢ (Catch c1 c2, Normal s) → (Catch c'' c2, s'')" by auto
        from hyp [OF this] 
        have "Γ⊢Catch c'' c2 ↓ s''".
        moreover
        from steps_Throw_impl_exec [OF rest]
        have "Γ⊢ ⟨c'',s''⟩ ⇒ Abrupt s'".
        moreover
        from rest obtain x where "s''=Normal x"
          by (cases s'')
             (auto dest: steps_Fault_prop steps_Abrupt_prop steps_Stuck_prop)
        ultimately show ?thesis
          by (fastforce elim: terminates_elim_cases)
      qed
    qed
  qed
next
  case (Await b c)       
  show ?case
  using Await.prems(2) by auto   
qed


lemma Await_finish:"⋀c2 s2 b c. Γ⊢ (Await b c, s1) → (c2, s2) ⟹ Γ⊢c2 ↓ s2"
by (metis Abrupt Fault Stuck final_termi step_await_final1 step_preserves_termination xstate.exhaust)

                                                            
lemma wf_implies_termi_reach:
assumes wf: "wf {(cfg2,cfg1). Γ ⊢ (c,s) →* cfg1 ∧ Γ ⊢ cfg1 → cfg2}"
shows "⋀c1 s1. ⟦Γ ⊢ (c,s) →* cfg1;  cfg1=(c1,s1)⟧⟹ Γ⊢c1↓s1"
using wf 
proof (induct cfg1,simp)
  fix c1 s1
  assume reach: "Γ⊢ (c, s) →* (c1, s1)"
  assume hyp_raw: "⋀y c2 s2.
           ⟦Γ⊢ (c1, s1) → (c2, s2); Γ⊢ (c, s) →* (c2, s2); y = (c2, s2)⟧
           ⟹ Γ⊢c2 ↓ s2"
  have hyp: "⋀c2 s2. Γ⊢ (c1, s1) → (c2, s2) ⟹ Γ⊢c2 ↓ s2"
    apply -
    apply (rule hyp_raw)
    apply   assumption
    using reach 
    apply  simp
    apply (rule refl)
    done
  
  show "Γ⊢c1 ↓ s1"
  proof (cases s1)  
    case (Normal s1')             
    with  wf_implies_termi_reach_step_case' [OF hyp [simplified Normal]] 
    show ?thesis
      by auto
  qed (auto intro: terminates.intros)
qed

theorem no_infinite_computation_impl_terminates:
  assumes not_inf: "¬ Γ⊢ (c, s) → …(∞)"
  shows "Γ⊢c↓s"
proof -
  from no_infinite_computation_implies_wf [OF not_inf]
  have wf: "wf {(c2, c1). Γ⊢(c, s) →* c1 ∧ Γ⊢c1 → c2}".
  show ?thesis
    by (rule wf_implies_termi_reach [OF wf]) auto
qed

corollary terminates_iff_no_infinite_computation: 
  "Γ⊢c↓s = (¬ Γ⊢ (c, s) → …(∞))"
  apply (rule)
  apply  (erule terminates_impl_no_infinite_computation)
  apply (erule no_infinite_computation_impl_terminates)
  done

(* ************************************************************************* *)
subsection {* Generalised Redexes *} 
(* ************************************************************************* *)

text {*
For an important lemma for the completeness proof of the Hoare-logic for
total correctness we need a generalisation of @{const "redex"} that not only
yield the redex itself but all the enclosing statements as well.
*}

primrec redexes:: "('s,'p,'f,'e)com ⇒ ('s,'p,'f,'e)com set"
where
"redexes Skip = {Skip}" |
"redexes (Basic f) = {Basic f}" |
"redexes (Spec r) = {Spec r}" |
"redexes (Seq c1 c2) = {Seq c1 c2} ∪ redexes c1" |
"redexes (Cond b c1 c2) = {Cond b c1 c2}" |
"redexes (While b c) = {While b c}" |
"redexes (Call p) = {Call p}" |
"redexes (DynCom d) = {DynCom d}" |
"redexes (Guard f b c) = {Guard f b c}" |
"redexes (Throw) = {Throw}" |
"redexes (Catch c1 c2) = {Catch c1 c2} ∪ redexes c1" |
"redexes (Await b c) = {Await b c}"

lemma root_in_redexes: "c ∈ redexes c"
  apply (induct c)
  apply auto
  done

lemma redex_in_redexes: "redex c ∈ redexes c"
  apply (induct c)
  apply auto
  done

lemma redex_redexes: "⋀c'. ⟦c' ∈ redexes c; redex c' = c'⟧ ⟹ redex c = c'" 
  apply (induct c)
  apply auto
  done

lemma step_redexes:
  shows "⋀r r'. ⟦Γ⊢(r,s) → (r',s'); r ∈ redexes c⟧ 
  ⟹ ∃c'. Γ⊢(c,s) → (c',s') ∧ r' ∈ redexes c'"
proof (induct c)
  case Skip thus ?case by (fastforce intro: step.intros elim: step_elim_cases)
next
  case Basic thus ?case by (fastforce intro: step.intros elim: step_elim_cases)
next
  case Spec thus ?case by (fastforce intro: step.intros elim: step_elim_cases)
next
  case (Seq c1 c2)
  have "r ∈ redexes (Seq c1 c2)" by fact
  hence r: "r = Seq c1 c2 ∨ r ∈ redexes c1"
    by simp
  have step_r: "Γ⊢ (r, s) → (r', s')" by fact
  from r show ?case
  proof 
    assume "r = Seq c1 c2"
    with step_r
    show ?case
      by (auto simp add: root_in_redexes)
  next
    assume r: "r ∈ redexes c1"
    from Seq.hyps (1) [OF step_r this] 
    obtain c' where 
      step_c1: "Γ⊢ (c1, s) → (c', s')" and
      r': "r' ∈ redexes c'"
      by blast
    from step.Seq [OF step_c1]
    have "Γ⊢ (Seq c1 c2, s) → (Seq c' c2, s')".
    with r'
    show ?case
      by auto
  qed
next
  case Cond 
  thus ?case 
    by (fastforce intro: step.intros elim: step_elim_cases simp add: root_in_redexes)
next
  case While 
  thus ?case 
    by (fastforce intro: step.intros elim: step_elim_cases simp add: root_in_redexes)
next
  case Call thus ?case 
    by (fastforce intro: step.intros elim: step_elim_cases simp add: root_in_redexes)
next
  case DynCom thus ?case 
    by (fastforce intro: step.intros elim: step_elim_cases simp add: root_in_redexes)
next
  case Guard thus ?case 
    by (fastforce intro: step.intros elim: step_elim_cases simp add: root_in_redexes)
next
  case Throw thus ?case 
    by (fastforce intro: step.intros elim: step_elim_cases simp add: root_in_redexes)
next
  case (Catch c1 c2)
  have "r ∈ redexes (Catch c1 c2)" by fact
  hence r: "r = Catch c1 c2 ∨ r ∈ redexes c1"
    by simp
  have step_r: "Γ⊢ (r, s) → (r', s')" by fact
  from r show ?case
  proof 
    assume "r = Catch c1 c2"
    with step_r
    show ?case
      by (auto simp add: root_in_redexes)
  next
    assume r: "r ∈ redexes c1"
    from Catch.hyps (1) [OF step_r this] 
    obtain c' where 
      step_c1: "Γ⊢ (c1, s) → (c', s')" and
      r': "r' ∈ redexes c'"
      by blast
    from step.Catch [OF step_c1]
    have "Γ⊢ (Catch c1 c2, s) → (Catch c' c2, s')".
    with r'
    show ?case
      by auto
  qed
next case (Await b c)  
     thus ?case 
     by (fastforce intro: step.intros elim: step_elim_cases simp add: root_in_redexes)
qed 

lemma steps_redexes:
  assumes steps: "Γ⊢ (r, s) →* (r', s')"
  shows "⋀c. r ∈ redexes c ⟹ ∃c'. Γ⊢(c,s) →* (c',s') ∧ r' ∈ redexes c'"
using steps 
proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans])
  case Refl
  then
  show "∃c'. Γ⊢ (c, s') →* (c', s') ∧ r' ∈ redexes c'"
    by auto
next
  case (Trans r s r'' s'')
  have "Γ⊢ (r, s) → (r'', s'')" "r ∈ redexes c" by fact+
  from step_redexes [OF this]
  obtain c' where
    step: "Γ⊢ (c, s) → (c', s'')" and
    r'': "r'' ∈ redexes c'"
    by blast
  note step
  also
  from Trans.hyps (3) [OF r'']
  obtain c'' where
    steps: "Γ⊢ (c', s'') →* (c'', s')" and
    r': "r' ∈ redexes c''"
    by blast
  note steps
  finally
  show ?case
    using r'
    by blast
qed



lemma steps_redexes':
  assumes steps: "Γ⊢ (r, s) →+ (r', s')"
  shows "⋀c. r ∈ redexes c ⟹ ∃c'. Γ⊢(c,s) →+ (c',s') ∧ r' ∈ redexes c'"
using steps 
proof (induct rule: tranclp_induct2 [consumes 1, case_names Step Trans])
  case (Step r' s' c') 
  have "Γ⊢ (r, s) → (r', s')" "r ∈ redexes c'" by fact+
  from step_redexes [OF this]
  show ?case
    by (blast intro: r_into_trancl)
next
  case (Trans r' s' r'' s'')
  from Trans obtain c' where
    steps: "Γ⊢ (c, s) →+ (c', s')" and
    r': "r' ∈ redexes c'"
    by blast
  note steps
  moreover
  have "Γ⊢ (r', s') → (r'', s'')" by fact
  from step_redexes [OF this r'] obtain c'' where
    step: "Γ⊢ (c', s') → (c'', s'')" and
    r'': "r'' ∈ redexes c''"
    by blast
  note step
  finally show ?case
    using r'' by blast
qed

lemma step_redexes_Seq:
  assumes step: "Γ⊢(r,s) → (r',s')"
  assumes Seq: "Seq r c2 ∈ redexes c"
  shows "∃c'. Γ⊢(c,s) → (c',s') ∧ Seq r' c2 ∈ redexes c'"
proof -
  from step.Seq [OF step]
  have "Γ⊢ (Seq r c2, s) → (Seq r' c2, s')".
  from step_redexes [OF this Seq] 
  show ?thesis .
qed

lemma steps_redexes_Seq:
  assumes steps: "Γ⊢ (r, s) →* (r', s')"
  shows "⋀c. Seq r c2 ∈ redexes c ⟹ 
              ∃c'. Γ⊢(c,s) →* (c',s') ∧ Seq r' c2 ∈ redexes c'"
using steps 
proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans])
  case Refl
  then show ?case
    by (auto)

next
  case (Trans r s r'' s'')
  have "Γ⊢ (r, s) → (r'', s'')" "Seq r c2 ∈ redexes c" by fact+
  from step_redexes_Seq [OF this]
  obtain c' where
    step: "Γ⊢ (c, s) → (c', s'')" and
    r'': "Seq r'' c2 ∈ redexes c'"
    by blast
  note step
  also
  from Trans.hyps (3) [OF r'']
  obtain c'' where
    steps: "Γ⊢ (c', s'') →* (c'', s')" and
    r': "Seq r' c2 ∈ redexes c''"
    by blast
  note steps
  finally
  show ?case
    using r'
    by blast
qed

lemma steps_redexes_Seq':
  assumes steps: "Γ⊢ (r, s) →+ (r', s')"
  shows "⋀c. Seq r c2 ∈ redexes c 
             ⟹ ∃c'. Γ⊢(c,s) →+ (c',s') ∧ Seq r' c2 ∈ redexes c'"
using steps 
proof (induct rule: tranclp_induct2 [consumes 1, case_names Step Trans])
  case (Step r' s' c') 
  have "Γ⊢ (r, s) → (r', s')" "Seq r c2 ∈ redexes c'" by fact+
  from step_redexes_Seq [OF this]
  show ?case
    by (blast intro: r_into_trancl)
next
  case (Trans r' s' r'' s'')
  from Trans obtain c' where
    steps: "Γ⊢ (c, s) →+ (c', s')" and
    r': "Seq r' c2 ∈ redexes c'"
    by blast
  note steps
  moreover
  have "Γ⊢ (r', s') → (r'', s'')" by fact
  from step_redexes_Seq [OF this r'] obtain c'' where
    step: "Γ⊢ (c', s') → (c'', s'')" and
    r'': "Seq r'' c2 ∈ redexes c''"
    by blast
  note step
  finally show ?case
    using r'' by blast
qed

lemma step_redexes_Catch:
  assumes step: "Γ⊢(r,s) → (r',s')"
  assumes Catch: "Catch r c2 ∈ redexes c"
  shows "∃c'. Γ⊢(c,s) → (c',s') ∧ Catch r' c2 ∈ redexes c'"
proof -
  from step.Catch [OF step]
  have "Γ⊢ (Catch r c2, s) → (Catch r' c2, s')".
  from step_redexes [OF this Catch] 
  show ?thesis .
qed

lemma steps_redexes_Catch:
  assumes steps: "Γ⊢ (r, s) →* (r', s')"
  shows "⋀c. Catch r c2 ∈ redexes c ⟹ 
              ∃c'. Γ⊢(c,s) →* (c',s') ∧ Catch r' c2 ∈ redexes c'"
using steps 
proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans])
  case Refl
  then show ?case
    by (auto)

next
  case (Trans r s r'' s'')
  have "Γ⊢ (r, s) → (r'', s'')" "Catch r c2 ∈ redexes c" by fact+
  from step_redexes_Catch [OF this]
  obtain c' where
    step: "Γ⊢ (c, s) → (c', s'')" and
    r'': "Catch r'' c2 ∈ redexes c'"
    by blast
  note step
  also
  from Trans.hyps (3) [OF r'']
  obtain c'' where
    steps: "Γ⊢ (c', s'') →* (c'', s')" and
    r': "Catch r' c2 ∈ redexes c''"
    by blast
  note steps
  finally
  show ?case
    using r'
    by blast
qed

lemma steps_redexes_Catch':
  assumes steps: "Γ⊢ (r, s) →+ (r', s')"
  shows "⋀c. Catch r c2 ∈ redexes c 
             ⟹ ∃c'. Γ⊢(c,s) →+ (c',s') ∧ Catch r' c2 ∈ redexes c'"
using steps 
proof (induct rule: tranclp_induct2 [consumes 1, case_names Step Trans])
  case (Step r' s' c') 
  have "Γ⊢ (r, s) → (r', s')" "Catch r c2 ∈ redexes c'" by fact+
  from step_redexes_Catch [OF this]
  show ?case
    by (blast intro: r_into_trancl)
next
  case (Trans r' s' r'' s'')
  from Trans obtain c' where
    steps: "Γ⊢ (c, s) →+ (c', s')" and
    r': "Catch r' c2 ∈ redexes c'"
    by blast
  note steps
  moreover
  have "Γ⊢ (r', s') → (r'', s'')" by fact
  from step_redexes_Catch [OF this r'] obtain c'' where
    step: "Γ⊢ (c', s') → (c'', s'')" and
    r'': "Catch r'' c2 ∈ redexes c''"
    by blast
  note step
  finally show ?case
    using r'' by blast
qed

lemma redexes_subset:"⋀c'. c' ∈ redexes c ⟹ redexes c' ⊆ redexes c"
  by (induct c) auto

lemma redexes_preserves_termination:
  assumes termi: "Γ⊢c↓s"
  shows "⋀c'. c' ∈ redexes c ⟹ Γ⊢c'↓s"  
using termi
by induct (auto intro: terminates.intros)

*)
end