Theory Common
theory Common
imports "../UNITY_Main"
begin
consts
ftime :: "nat=>nat"
gtime :: "nat=>nat"
axiomatization where
fmono: "m ≤ n ==> ftime m ≤ ftime n" and
gmono: "m ≤ n ==> gtime m ≤ gtime n" and
fasc: "m ≤ ftime n" and
gasc: "m ≤ gtime n"
definition common :: "nat set" where
"common == {n. ftime n = n & gtime n = n}"
definition maxfg :: "nat => nat set" where
"maxfg m == {t. t ≤ max (ftime m) (gtime m)}"
lemma common_stable:
"[| ∀m. F ∈ {m} Co (maxfg m); n ∈ common |]
==> F ∈ Stable (atMost n)"
apply (drule_tac M = "{t. t ≤ n}" in Elimination_sing)
apply (simp add: atMost_def Stable_def common_def maxfg_def le_max_iff_disj)
apply (erule Constrains_weaken_R)
apply (blast intro: order_eq_refl le_trans dest: fmono gmono)
done
lemma common_safety:
"[| Init F ⊆ atMost n;
∀m. F ∈ {m} Co (maxfg m); n ∈ common |]
==> F ∈ Always (atMost n)"
by (simp add: AlwaysI common_stable)
lemma "SKIP ∈ {m} co (maxfg m)"
by (simp add: constrains_def maxfg_def le_max_iff_disj fasc)
lemma "mk_total_program
(UNIV, {range(%t.(t,ftime t)), range(%t.(t,gtime t))}, UNIV)
∈ {m} co (maxfg m)"
apply (simp add: mk_total_program_def)
apply (simp add: constrains_def maxfg_def le_max_iff_disj fasc)
done
lemma "mk_total_program (UNIV, {range(%t.(t, max (ftime t) (gtime t)))}, UNIV)
∈ {m} co (maxfg m)"
apply (simp add: mk_total_program_def)
apply (simp add: constrains_def maxfg_def gasc max.absorb2)
done
lemma "mk_total_program
(UNIV, { {(t, Suc t) | t. t < max (ftime t) (gtime t)} }, UNIV)
∈ {m} co (maxfg m)"
apply (simp add: mk_total_program_def)
apply (simp add: constrains_def maxfg_def gasc max.absorb2)
done
lemma leadsTo_common_lemma:
assumes "∀m. F ∈ {m} Co (maxfg m)"
and "∀m ∈ lessThan n. F ∈ {m} LeadsTo (greaterThan m)"
and "n ∈ common"
shows "F ∈ (atMost n) LeadsTo common"
proof (rule LeadsTo_weaken_R)
show "F ∈ {..n} LeadsTo {..n} ∩ id -` {n..} ∪ common"
proof (induct rule: GreaterThan_bounded_induct [of n _ _ id])
case 1
from assms have "∀m∈{..<n}. F ∈ {..n} ∩ {m} LeadsTo {..n} ∩ {m<..} ∪ common"
by (blast dest: PSP_Stable2 intro: common_stable LeadsTo_weaken_R)
then show ?case by simp
qed
next
from ‹n ∈ common›
show "{..n} ∩ id -` {n..} ∪ common ⊆ common"
by (simp add: atMost_Int_atLeast)
qed
lemma leadsTo_common:
"[| ∀m. F ∈ {m} Co (maxfg m);
∀m ∈ -common. F ∈ {m} LeadsTo (greaterThan m);
n ∈ common |]
==> F ∈ (atMost (LEAST n. n ∈ common)) LeadsTo common"
apply (rule leadsTo_common_lemma)
apply (simp_all (no_asm_simp))
apply (erule_tac [2] LeastI)
apply (blast dest!: not_less_Least)
done
end