Theory Kripke
section ‹Kripke Structures and CTL›
theory Kripke
imports Main
begin
definition
Kripke :: "['s set,
's set,
('s * 's) set,
('s ⇀ 'a set)]
=> bool" where
"Kripke S S0 R L =
(S0 ⊆ S ∧
R <= S × S ∧
(Domain R) = S ∧
(dom L) = S)"
lemma Kripke_EmptySet:
"({@x. True}, {@x. True},{(@x. True, @x. True)}, Map.empty(@x. True ↦ {@x. True})) ∈
{(S,S0,R,L) | S S0 R L. Kripke S S0 R L}"
by (unfold Kripke_def Domain_unfold, auto)
definition
"kripke =
{(S,S0,T,L) |
(S::('s set))
(S0::('s set))
(T::(('s * 's) set))
(L::('s ⇀ ('a set))).
Kripke S S0 T L}"
typedef ('s,'a) kripke =
"kripke :: ('s set * 's set * ('s * 's) set * ('s ⇀ 'a set)) set"
unfolding kripke_def
apply (rule exI)
apply (rule Kripke_EmptySet)
done
definition
Statuses :: "('s,'a) kripke => 's set" where
"Statuses = fst o Rep_kripke"
definition
InitStatuses :: "('s,'a) kripke => 's set" where
"InitStatuses == fst o snd o Rep_kripke"
definition
StepRel :: "('s,'a) kripke => ('s * 's) set" where
"StepRel == fst o snd o snd o Rep_kripke"
definition
LabelFun :: "('s,'a) kripke => ('s ⇀ 'a set)" where
"LabelFun == snd o snd o snd o Rep_kripke"
definition
Paths :: "[('s,'a) kripke, 's] =>
(nat => 's) set" where
"Paths M S == { p . S = p (0::nat) ∧ (∀i. (p i, p (i+1)) ∈ (StepRel M))}"