Theory Renaming_Non_Destructive
section ‹Non Destructiveness of Renaming›
theory Renaming_Non_Destructive
imports Process_Restriction_Space "HOL-CSPM.CSPM_Laws"
begin
subsection ‹Equality›
lemma restriction_process⇩p⇩t⇩i⇩c⇩k_Renaming:
‹Renaming P f g ↓ n = Renaming (P ↓ n) f g› (is ‹?lhs = ?rhs›)
proof (rule Process_eq_optimizedI)
show ‹t ∈ 𝒟 ?lhs ⟹ t ∈ 𝒟 ?rhs› for t
by (auto simp add: Renaming_projs D_restriction_process⇩p⇩t⇩i⇩c⇩k)
(metis append.right_neutral front_tickFree_Nil map_event⇩p⇩t⇩i⇩c⇩k_tickFree,
use front_tickFree_append in blast)
next
show ‹t ∈ 𝒟 ?rhs ⟹ t ∈ 𝒟 ?lhs› for t
by (auto simp add: Renaming_projs D_restriction_process⇩p⇩t⇩i⇩c⇩k
front_tickFree_append map_event⇩p⇩t⇩i⇩c⇩k_tickFree)
next
fix t X assume ‹(t, X) ∈ ℱ ?lhs› ‹t ∉ 𝒟 ?lhs›
then obtain u where ‹(u, map_event⇩p⇩t⇩i⇩c⇩k f g -` X) ∈ ℱ P› ‹t = map (map_event⇩p⇩t⇩i⇩c⇩k f g) u›
by (simp add: Renaming_projs restriction_process⇩p⇩t⇩i⇩c⇩k_projs) blast
thus ‹(t, X) ∈ ℱ ?rhs› by (auto simp add: F_Renaming F_restriction_process⇩p⇩t⇩i⇩c⇩k)
next
fix t X assume ‹(t, X) ∈ ℱ ?rhs› ‹t ∉ 𝒟 ?rhs›
then obtain u where ‹(u, map_event⇩p⇩t⇩i⇩c⇩k f g -` X) ∈ ℱ (P ↓ n)› ‹t = map (map_event⇩p⇩t⇩i⇩c⇩k f g) u›
unfolding Renaming_projs by blast
from ‹(u, map_event⇩p⇩t⇩i⇩c⇩k f g -` X) ∈ ℱ (P ↓ n)›
consider ‹u ∈ 𝒟 (P ↓ n)› | ‹(u, map_event⇩p⇩t⇩i⇩c⇩k f g -` X) ∈ ℱ P›
unfolding restriction_process⇩p⇩t⇩i⇩c⇩k_projs by blast
thus ‹(t, X) ∈ ℱ ?lhs›
proof cases
assume ‹u ∈ 𝒟 (P ↓ n)›
hence ‹t ∈ 𝒟 ?rhs›
proof (elim D_restriction_process⇩p⇩t⇩i⇩c⇩kE)
from ‹t = map (map_event⇩p⇩t⇩i⇩c⇩k f g) u› show ‹u ∈ 𝒟 P ⟹ t ∈ 𝒟 ?rhs›
by (cases ‹tF u›, simp_all add: D_Renaming D_restriction_process⇩p⇩t⇩i⇩c⇩k)
(use front_tickFree_Nil in blast,
metis D_imp_front_tickFree butlast_snoc div_butlast_when_non_tickFree_iff
front_tickFree_iff_tickFree_butlast front_tickFree_single map_append
map_event⇩p⇩t⇩i⇩c⇩k_front_tickFree nonTickFree_n_frontTickFree)
next
show ‹⟦u = v @ w; v ∈ 𝒯 P; length v = n; tF v; ftF w⟧ ⟹ t ∈ 𝒟 ?rhs› for v w
by (simp add: D_Renaming D_restriction_process⇩p⇩t⇩i⇩c⇩k ‹t = map (map_event⇩p⇩t⇩i⇩c⇩k f g) u›)
(use front_tickFree_Nil map_event⇩p⇩t⇩i⇩c⇩k_front_tickFree in blast)
qed
with ‹t ∉ 𝒟 ?rhs› have False ..
thus ‹(t, X) ∈ ℱ ?lhs› ..
next
show ‹(u, map_event⇩p⇩t⇩i⇩c⇩k f g -` X) ∈ ℱ P ⟹ (t, X) ∈ ℱ (Renaming P f g ↓ n)›
by (auto simp add: F_restriction_process⇩p⇩t⇩i⇩c⇩k F_Renaming ‹t = map (map_event⇩p⇩t⇩i⇩c⇩k f g) u›)
qed
qed
subsection ‹Non Destructiveness›
lemma Renaming_non_destructive [simp] :
‹non_destructive (λP. Renaming P f g)›
by (auto intro: non_destructiveI simp add: restriction_process⇩p⇩t⇩i⇩c⇩k_Renaming)
end