src/Doc/Tutorial/CTL/Base.thy
changeset 69505 cc2d676d5395
parent 67406 23307fd33906
child 69597 ff784d5a5bfb
equal deleted inserted replaced
69504:bda7527ccf05 69505:cc2d676d5395
    10 (specifications) @{cite "ClarkeGP-book" and "Huth-Ryan-book"}. Its foundations are set theoretic
    10 (specifications) @{cite "ClarkeGP-book" and "Huth-Ryan-book"}. Its foundations are set theoretic
    11 and this section will explore them in HOL\@. This is done in two steps.  First
    11 and this section will explore them in HOL\@. This is done in two steps.  First
    12 we consider a simple modal logic called propositional dynamic
    12 we consider a simple modal logic called propositional dynamic
    13 logic (PDL)\@.  We then proceed to the temporal logic CTL, which is
    13 logic (PDL)\@.  We then proceed to the temporal logic CTL, which is
    14 used in many real
    14 used in many real
    15 model checkers. In each case we give both a traditional semantics (@{text \<Turnstile>}) and a
    15 model checkers. In each case we give both a traditional semantics (\<open>\<Turnstile>\<close>) and a
    16 recursive function @{term mc} that maps a formula into the set of all states of
    16 recursive function @{term mc} that maps a formula into the set of all states of
    17 the system where the formula is valid. If the system has a finite number of
    17 the system where the formula is valid. If the system has a finite number of
    18 states, @{term mc} is directly executable: it is a model checker, albeit an
    18 states, @{term mc} is directly executable: it is a model checker, albeit an
    19 inefficient one. The main proof obligation is to show that the semantics
    19 inefficient one. The main proof obligation is to show that the semantics
    20 and the model checker agree.
    20 and the model checker agree.