src/HOL/ex/CTL.thy
author wenzelm
Mon, 24 Oct 2016 14:05:22 +0200
changeset 64371 213cf4215b40
parent 63055 ae0ca486bd3f
child 69597 ff784d5a5bfb
permissions -rw-r--r--
more robust;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
     1
(*  Title:      HOL/ex/CTL.thy
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
     2
    Author:     Gertrud Bauer
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
     3
*)
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
     4
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     5
section \<open>CTL formulae\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
     6
46685
866a798d051c tuned proofs;
wenzelm
parents: 46008
diff changeset
     7
theory CTL
866a798d051c tuned proofs;
wenzelm
parents: 46008
diff changeset
     8
imports Main
866a798d051c tuned proofs;
wenzelm
parents: 46008
diff changeset
     9
begin
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    10
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    11
text \<open>
61934
wenzelm
parents: 61933
diff changeset
    12
  We formalize basic concepts of Computational Tree Logic (CTL) @{cite
wenzelm
parents: 61933
diff changeset
    13
  "McMillan-PhDThesis" and "McMillan-LectureNotes"} within the simply-typed
wenzelm
parents: 61933
diff changeset
    14
  set theory of HOL.
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    15
61934
wenzelm
parents: 61933
diff changeset
    16
  By using the common technique of ``shallow embedding'', a CTL formula is
wenzelm
parents: 61933
diff changeset
    17
  identified with the corresponding set of states where it holds.
wenzelm
parents: 61933
diff changeset
    18
  Consequently, CTL operations such as negation, conjunction, disjunction
wenzelm
parents: 61933
diff changeset
    19
  simply become complement, intersection, union of sets. We only require a
wenzelm
parents: 61933
diff changeset
    20
  separate operation for implication, as point-wise inclusion is usually not
wenzelm
parents: 61933
diff changeset
    21
  encountered in plain set-theory.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    22
\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    23
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    24
lemmas [intro!] = Int_greatest Un_upper2 Un_upper1 Int_lower1 Int_lower2
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    25
42463
f270e3e18be5 modernized specifications;
wenzelm
parents: 41460
diff changeset
    26
type_synonym 'a ctl = "'a set"
20807
wenzelm
parents: 17388
diff changeset
    27
63054
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61934
diff changeset
    28
definition imp :: "'a ctl \<Rightarrow> 'a ctl \<Rightarrow> 'a ctl"  (infixr "\<rightarrow>" 75)
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61934
diff changeset
    29
  where "p \<rightarrow> q = - p \<union> q"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    30
20807
wenzelm
parents: 17388
diff changeset
    31
lemma [intro!]: "p \<inter> p \<rightarrow> q \<subseteq> q" unfolding imp_def by auto
wenzelm
parents: 17388
diff changeset
    32
lemma [intro!]: "p \<subseteq> (q \<rightarrow> p)" unfolding imp_def by rule
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    33
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    34
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    35
text \<open>
61934
wenzelm
parents: 61933
diff changeset
    36
  \<^smallskip>
wenzelm
parents: 61933
diff changeset
    37
  The CTL path operators are more interesting; they are based on an arbitrary,
wenzelm
parents: 61933
diff changeset
    38
  but fixed model \<open>\<M>\<close>, which is simply a transition relation over states
wenzelm
parents: 61933
diff changeset
    39
  @{typ 'a}.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    40
\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    41
20807
wenzelm
parents: 17388
diff changeset
    42
axiomatization \<M> :: "('a \<times> 'a) set"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    43
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    44
text \<open>
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    45
  The operators \<open>\<^bold>E\<^bold>X\<close>, \<open>\<^bold>E\<^bold>F\<close>, \<open>\<^bold>E\<^bold>G\<close> are taken as primitives, while \<open>\<^bold>A\<^bold>X\<close>,
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    46
  \<open>\<^bold>A\<^bold>F\<close>, \<open>\<^bold>A\<^bold>G\<close> are defined as derived ones. The formula \<open>\<^bold>E\<^bold>X p\<close> holds in a
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    47
  state \<open>s\<close>, iff there is a successor state \<open>s'\<close> (with respect to the model
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    48
  \<open>\<M>\<close>), such that \<open>p\<close> holds in \<open>s'\<close>. The formula \<open>\<^bold>E\<^bold>F p\<close> holds in a state
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    49
  \<open>s\<close>, iff there is a path in \<open>\<M>\<close>, starting from \<open>s\<close>, such that there exists a
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    50
  state \<open>s'\<close> on the path, such that \<open>p\<close> holds in \<open>s'\<close>. The formula \<open>\<^bold>E\<^bold>G p\<close>
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    51
  holds in a state \<open>s\<close>, iff there is a path, starting from \<open>s\<close>, such that for
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    52
  all states \<open>s'\<close> on the path, \<open>p\<close> holds in \<open>s'\<close>. It is easy to see that \<open>\<^bold>E\<^bold>F
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    53
  p\<close> and \<open>\<^bold>E\<^bold>G p\<close> may be expressed using least and greatest fixed points
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    54
  @{cite "McMillan-PhDThesis"}.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    55
\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    56
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    57
definition EX  ("\<^bold>E\<^bold>X _" [80] 90)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    58
  where [simp]: "\<^bold>E\<^bold>X p = {s. \<exists>s'. (s, s') \<in> \<M> \<and> s' \<in> p}"
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    59
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    60
definition EF ("\<^bold>E\<^bold>F _" [80] 90)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    61
  where [simp]: "\<^bold>E\<^bold>F p = lfp (\<lambda>s. p \<union> \<^bold>E\<^bold>X s)"
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    62
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    63
definition EG ("\<^bold>E\<^bold>G _" [80] 90)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    64
  where [simp]: "\<^bold>E\<^bold>G p = gfp (\<lambda>s. p \<inter> \<^bold>E\<^bold>X s)"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    65
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    66
text \<open>
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    67
  \<open>\<^bold>A\<^bold>X\<close>, \<open>\<^bold>A\<^bold>F\<close> and \<open>\<^bold>A\<^bold>G\<close> are now defined dually in terms of \<open>\<^bold>E\<^bold>X\<close>,
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    68
  \<open>\<^bold>E\<^bold>F\<close> and \<open>\<^bold>E\<^bold>G\<close>.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    69
\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    70
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    71
definition AX  ("\<^bold>A\<^bold>X _" [80] 90)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    72
  where [simp]: "\<^bold>A\<^bold>X p = - \<^bold>E\<^bold>X - p"
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    73
definition AF  ("\<^bold>A\<^bold>F _" [80] 90)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    74
  where [simp]: "\<^bold>A\<^bold>F p = - \<^bold>E\<^bold>G - p"
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    75
definition AG  ("\<^bold>A\<^bold>G _" [80] 90)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
    76
  where [simp]: "\<^bold>A\<^bold>G p = - \<^bold>E\<^bold>F - p"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    77
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    78
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    79
subsection \<open>Basic fixed point properties\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    80
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    81
text \<open>
61934
wenzelm
parents: 61933
diff changeset
    82
  First of all, we use the de-Morgan property of fixed points.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    83
\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    84
21026
3b2821e0d541 Adapted to changes in FixedPoint theory.
berghofe
parents: 20807
diff changeset
    85
lemma lfp_gfp: "lfp f = - gfp (\<lambda>s::'a set. - (f (- s)))"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    86
proof
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    87
  show "lfp f \<subseteq> - gfp (\<lambda>s. - f (- s))"
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    88
  proof
63054
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61934
diff changeset
    89
    show "x \<in> - gfp (\<lambda>s. - f (- s))" if l: "x \<in> lfp f" for x
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    90
    proof
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    91
      assume "x \<in> gfp (\<lambda>s. - f (- s))"
21026
3b2821e0d541 Adapted to changes in FixedPoint theory.
berghofe
parents: 20807
diff changeset
    92
      then obtain u where "x \<in> u" and "u \<subseteq> - f (- u)"
32587
caa5ada96a00 Inter and Union are mere abbreviations for Inf and Sup
haftmann
parents: 26813
diff changeset
    93
        by (auto simp add: gfp_def)
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    94
      then have "f (- u) \<subseteq> - u" by auto
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    95
      then have "lfp f \<subseteq> - u" by (rule lfp_lowerbound)
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    96
      from l and this have "x \<notin> u" by auto
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    97
      with \<open>x \<in> u\<close> show False by contradiction
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    98
    qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
    99
  qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   100
  show "- gfp (\<lambda>s. - f (- s)) \<subseteq> lfp f"
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   101
  proof (rule lfp_greatest)
63054
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61934
diff changeset
   102
    fix u
1b237d147cc4 misc tuning and modernization;
wenzelm
parents: 61934
diff changeset
   103
    assume "f u \<subseteq> u"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   104
    then have "- u \<subseteq> - f u" by auto
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   105
    then have "- u \<subseteq> - f (- (- u))" by simp
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   106
    then have "- u \<subseteq> gfp (\<lambda>s. - f (- s))" by (rule gfp_upperbound)
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   107
    then show "- gfp (\<lambda>s. - f (- s)) \<subseteq> u" by auto
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   108
  qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   109
qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   110
21026
3b2821e0d541 Adapted to changes in FixedPoint theory.
berghofe
parents: 20807
diff changeset
   111
lemma lfp_gfp': "- lfp f = gfp (\<lambda>s::'a set. - (f (- s)))"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   112
  by (simp add: lfp_gfp)
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   113
21026
3b2821e0d541 Adapted to changes in FixedPoint theory.
berghofe
parents: 20807
diff changeset
   114
lemma gfp_lfp': "- gfp f = lfp (\<lambda>s::'a set. - (f (- s)))"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   115
  by (simp add: lfp_gfp)
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   116
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   117
text \<open>
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   118
  In order to give dual fixed point representations of @{term "\<^bold>A\<^bold>F p"} and
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   119
  @{term "\<^bold>A\<^bold>G p"}:
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   120
\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   121
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   122
lemma AF_lfp: "\<^bold>A\<^bold>F p = lfp (\<lambda>s. p \<union> \<^bold>A\<^bold>X s)"
61934
wenzelm
parents: 61933
diff changeset
   123
  by (simp add: lfp_gfp)
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   124
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   125
lemma AG_gfp: "\<^bold>A\<^bold>G p = gfp (\<lambda>s. p \<inter> \<^bold>A\<^bold>X s)"
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   126
  by (simp add: lfp_gfp)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   127
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   128
lemma EF_fp: "\<^bold>E\<^bold>F p = p \<union> \<^bold>E\<^bold>X \<^bold>E\<^bold>F p"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   129
proof -
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   130
  have "mono (\<lambda>s. p \<union> \<^bold>E\<^bold>X s)" by rule auto
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   131
  then show ?thesis by (simp only: EF_def) (rule lfp_unfold)
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   132
qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   133
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   134
lemma AF_fp: "\<^bold>A\<^bold>F p = p \<union> \<^bold>A\<^bold>X \<^bold>A\<^bold>F p"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   135
proof -
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   136
  have "mono (\<lambda>s. p \<union> \<^bold>A\<^bold>X s)" by rule auto
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   137
  then show ?thesis by (simp only: AF_lfp) (rule lfp_unfold)
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   138
qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   139
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   140
lemma EG_fp: "\<^bold>E\<^bold>G p = p \<inter> \<^bold>E\<^bold>X \<^bold>E\<^bold>G p"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   141
proof -
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   142
  have "mono (\<lambda>s. p \<inter> \<^bold>E\<^bold>X s)" by rule auto
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   143
  then show ?thesis by (simp only: EG_def) (rule gfp_unfold)
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   144
qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   145
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   146
text \<open>
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   147
  From the greatest fixed point definition of @{term "\<^bold>A\<^bold>G p"}, we derive as
61934
wenzelm
parents: 61933
diff changeset
   148
  a consequence of the Knaster-Tarski theorem on the one hand that @{term
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   149
  "\<^bold>A\<^bold>G p"} is a fixed point of the monotonic function
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   150
  @{term "\<lambda>s. p \<inter> \<^bold>A\<^bold>X s"}.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   151
\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   152
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   153
lemma AG_fp: "\<^bold>A\<^bold>G p = p \<inter> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   154
proof -
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   155
  have "mono (\<lambda>s. p \<inter> \<^bold>A\<^bold>X s)" by rule auto
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   156
  then show ?thesis by (simp only: AG_gfp) (rule gfp_unfold)
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   157
qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   158
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   159
text \<open>
61934
wenzelm
parents: 61933
diff changeset
   160
  This fact may be split up into two inequalities (merely using transitivity
wenzelm
parents: 61933
diff changeset
   161
  of \<open>\<subseteq>\<close>, which is an instance of the overloaded \<open>\<le>\<close> in Isabelle/HOL).
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   162
\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   163
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   164
lemma AG_fp_1: "\<^bold>A\<^bold>G p \<subseteq> p"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   165
proof -
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   166
  note AG_fp also have "p \<inter> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p \<subseteq> p" by auto
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   167
  finally show ?thesis .
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   168
qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   169
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   170
lemma AG_fp_2: "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   171
proof -
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   172
  note AG_fp also have "p \<inter> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" by auto
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   173
  finally show ?thesis .
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   174
qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   175
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   176
text \<open>
61934
wenzelm
parents: 61933
diff changeset
   177
  On the other hand, we have from the Knaster-Tarski fixed point theorem that
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   178
  any other post-fixed point of @{term "\<lambda>s. p \<inter> \<^bold>A\<^bold>X s"} is smaller than
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   179
  @{term "\<^bold>A\<^bold>G p"}. A post-fixed point is a set of states \<open>q\<close> such that @{term
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   180
  "q \<subseteq> p \<inter> \<^bold>A\<^bold>X q"}. This leads to the following co-induction principle for
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   181
  @{term "\<^bold>A\<^bold>G p"}.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   182
\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   183
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   184
lemma AG_I: "q \<subseteq> p \<inter> \<^bold>A\<^bold>X q \<Longrightarrow> q \<subseteq> \<^bold>A\<^bold>G p"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   185
  by (simp only: AG_gfp) (rule gfp_upperbound)
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   186
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   187
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   188
subsection \<open>The tree induction principle \label{sec:calc-ctl-tree-induct}\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   189
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   190
text \<open>
61934
wenzelm
parents: 61933
diff changeset
   191
  With the most basic facts available, we are now able to establish a few more
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   192
  interesting results, leading to the \<^emph>\<open>tree induction\<close> principle for \<open>\<^bold>A\<^bold>G\<close>
61934
wenzelm
parents: 61933
diff changeset
   193
  (see below). We will use some elementary monotonicity and distributivity
wenzelm
parents: 61933
diff changeset
   194
  rules.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   195
\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   196
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   197
lemma AX_int: "\<^bold>A\<^bold>X (p \<inter> q) = \<^bold>A\<^bold>X p \<inter> \<^bold>A\<^bold>X q" by auto
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   198
lemma AX_mono: "p \<subseteq> q \<Longrightarrow> \<^bold>A\<^bold>X p \<subseteq> \<^bold>A\<^bold>X q" by auto
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   199
lemma AG_mono: "p \<subseteq> q \<Longrightarrow> \<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>G q"
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   200
  by (simp only: AG_gfp, rule gfp_mono) auto
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   201
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   202
text \<open>
61934
wenzelm
parents: 61933
diff changeset
   203
  The formula @{term "AG p"} implies @{term "AX p"} (we use substitution of
wenzelm
parents: 61933
diff changeset
   204
  \<open>\<subseteq>\<close> with monotonicity).
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   205
\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   206
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   207
lemma AG_AX: "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>X p"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   208
proof -
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   209
  have "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" by (rule AG_fp_2)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   210
  also have "\<^bold>A\<^bold>G p \<subseteq> p" by (rule AG_fp_1)
61934
wenzelm
parents: 61933
diff changeset
   211
  moreover note AX_mono
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   212
  finally show ?thesis .
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   213
qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   214
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   215
text \<open>
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   216
  Furthermore we show idempotency of the \<open>\<^bold>A\<^bold>G\<close> operator. The proof is a good
61934
wenzelm
parents: 61933
diff changeset
   217
  example of how accumulated facts may get used to feed a single rule step.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   218
\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   219
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   220
lemma AG_AG: "\<^bold>A\<^bold>G \<^bold>A\<^bold>G p = \<^bold>A\<^bold>G p"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   221
proof
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   222
  show "\<^bold>A\<^bold>G \<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>G p" by (rule AG_fp_1)
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   223
next
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   224
  show "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>G \<^bold>A\<^bold>G p"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   225
  proof (rule AG_I)
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   226
    have "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>G p" ..
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   227
    moreover have "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" by (rule AG_fp_2)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   228
    ultimately show "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>G p \<inter> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" ..
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   229
  qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   230
qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   231
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   232
text \<open>
61934
wenzelm
parents: 61933
diff changeset
   233
  \<^smallskip>
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   234
  We now give an alternative characterization of the \<open>\<^bold>A\<^bold>G\<close> operator, which
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   235
  describes the \<open>\<^bold>A\<^bold>G\<close> operator in an ``operational'' way by tree induction:
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   236
  In a state holds @{term "AG p"} iff in that state holds \<open>p\<close>, and in all
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   237
  reachable states \<open>s\<close> follows from the fact that \<open>p\<close> holds in \<open>s\<close>, that \<open>p\<close>
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   238
  also holds in all successor states of \<open>s\<close>. We use the co-induction principle
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   239
  @{thm [source] AG_I} to establish this in a purely algebraic manner.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   240
\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   241
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   242
theorem AG_induct: "p \<inter> \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p) = \<^bold>A\<^bold>G p"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   243
proof
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   244
  show "p \<inter> \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p) \<subseteq> \<^bold>A\<^bold>G p"  (is "?lhs \<subseteq> _")
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   245
  proof (rule AG_I)
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   246
    show "?lhs \<subseteq> p \<inter> \<^bold>A\<^bold>X ?lhs"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   247
    proof
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   248
      show "?lhs \<subseteq> p" ..
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   249
      show "?lhs \<subseteq> \<^bold>A\<^bold>X ?lhs"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   250
      proof -
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32587
diff changeset
   251
        {
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   252
          have "\<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p) \<subseteq> p \<rightarrow> \<^bold>A\<^bold>X p" by (rule AG_fp_1)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   253
          also have "p \<inter> p \<rightarrow> \<^bold>A\<^bold>X p \<subseteq> \<^bold>A\<^bold>X p" ..
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   254
          finally have "?lhs \<subseteq> \<^bold>A\<^bold>X p" by auto
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   255
        }
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32587
diff changeset
   256
        moreover
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32587
diff changeset
   257
        {
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   258
          have "p \<inter> \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p) \<subseteq> \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p)" ..
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   259
          also have "\<dots> \<subseteq> \<^bold>A\<^bold>X \<dots>" by (rule AG_fp_2)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   260
          finally have "?lhs \<subseteq> \<^bold>A\<^bold>X \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p)" .
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   261
        }
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   262
        ultimately have "?lhs \<subseteq> \<^bold>A\<^bold>X p \<inter> \<^bold>A\<^bold>X \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p)" ..
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   263
        also have "\<dots> = \<^bold>A\<^bold>X ?lhs" by (simp only: AX_int)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32587
diff changeset
   264
        finally show ?thesis .
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   265
      qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   266
    qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   267
  qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   268
next
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   269
  show "\<^bold>A\<^bold>G p \<subseteq> p \<inter> \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p)"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   270
  proof
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   271
    show "\<^bold>A\<^bold>G p \<subseteq> p" by (rule AG_fp_1)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   272
    show "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p)"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   273
    proof -
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   274
      have "\<^bold>A\<^bold>G p = \<^bold>A\<^bold>G \<^bold>A\<^bold>G p" by (simp only: AG_AG)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   275
      also have "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>X p" by (rule AG_AX) moreover note AG_mono
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   276
      also have "\<^bold>A\<^bold>X p \<subseteq> (p \<rightarrow> \<^bold>A\<^bold>X p)" .. moreover note AG_mono
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   277
      finally show ?thesis .
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   278
    qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   279
  qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   280
qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   281
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   282
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   283
subsection \<open>An application of tree induction \label{sec:calc-ctl-commute}\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   284
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   285
text \<open>
61934
wenzelm
parents: 61933
diff changeset
   286
  Further interesting properties of CTL expressions may be demonstrated with
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   287
  the help of tree induction; here we show that \<open>\<^bold>A\<^bold>X\<close> and \<open>\<^bold>A\<^bold>G\<close> commute.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   288
\<close>
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   289
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   290
theorem AG_AX_commute: "\<^bold>A\<^bold>G \<^bold>A\<^bold>X p = \<^bold>A\<^bold>X \<^bold>A\<^bold>G p"
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   291
proof -
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   292
  have "\<^bold>A\<^bold>G \<^bold>A\<^bold>X p = \<^bold>A\<^bold>X p \<inter> \<^bold>A\<^bold>X \<^bold>A\<^bold>G \<^bold>A\<^bold>X p" by (rule AG_fp)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   293
  also have "\<dots> = \<^bold>A\<^bold>X (p \<inter> \<^bold>A\<^bold>G \<^bold>A\<^bold>X p)" by (simp only: AX_int)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   294
  also have "p \<inter> \<^bold>A\<^bold>G \<^bold>A\<^bold>X p = \<^bold>A\<^bold>G p"  (is "?lhs = _")
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   295
  proof
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   296
    have "\<^bold>A\<^bold>X p \<subseteq> p \<rightarrow> \<^bold>A\<^bold>X p" ..
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   297
    also have "p \<inter> \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p) = \<^bold>A\<^bold>G p" by (rule AG_induct)
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   298
    also note Int_mono AG_mono
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   299
    ultimately show "?lhs \<subseteq> \<^bold>A\<^bold>G p" by fast
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   300
  next
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   301
    have "\<^bold>A\<^bold>G p \<subseteq> p" by (rule AG_fp_1)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   302
    moreover
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   303
    {
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   304
      have "\<^bold>A\<^bold>G p = \<^bold>A\<^bold>G \<^bold>A\<^bold>G p" by (simp only: AG_AG)
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   305
      also have "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>X p" by (rule AG_AX)
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   306
      also note AG_mono
63055
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   307
      ultimately have "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>G \<^bold>A\<^bold>X p" .
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   308
    }
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   309
    ultimately show "\<^bold>A\<^bold>G p \<subseteq> ?lhs" ..
ae0ca486bd3f tuned notation;
wenzelm
parents: 63054
diff changeset
   310
  qed
15871
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   311
  finally show ?thesis .
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   312
qed
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   313
e524119dbf19 *** empty log message ***
bauerg
parents:
diff changeset
   314
end