src/HOL/IMP/Abs_State.thy
author nipkow
Thu, 13 Sep 2012 10:28:48 +0200
changeset 49344 ce1ccb78ecda
parent 47619 0d3e95375bb7
child 49353 023be49d7fb8
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
     2
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
     3
theory Abs_State
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
     4
imports Abs_Int0
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
     5
begin
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
     6
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
     7
subsubsection "Welltypedness: wt"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
     8
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
     9
instantiation com :: vars
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    10
begin
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    11
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    12
fun vars_com :: "com \<Rightarrow> vname set" where
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    13
"vars com.SKIP = {}" |
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    14
"vars (x::=e) = {x} \<union> vars e" |
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    15
"vars (c1;c2) = vars c1 \<union> vars c2" |
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    16
"vars (IF b THEN c1 ELSE c2) = vars b \<union> vars c1 \<union> vars c2" |
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    17
"vars (WHILE b DO c) = vars b \<union> vars c"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    18
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    19
instance ..
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    20
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    21
end
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    22
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    23
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    24
lemma finite_avars: "finite(vars(a::aexp))"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    25
by(induction a) simp_all
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    26
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    27
lemma finite_bvars: "finite(vars(b::bexp))"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    28
by(induction b) (simp_all add: finite_avars)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    29
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    30
lemma finite_cvars: "finite(vars(c::com))"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    31
by(induction c) (simp_all add: finite_avars finite_bvars)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    32
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    33
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    34
class wt =
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    35
fixes wt :: "'a \<Rightarrow> vname set \<Rightarrow> bool"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    36
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    37
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    38
instantiation acom :: (wt)wt
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    39
begin
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    40
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    41
definition wt_acom where
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    42
"wt C X = (vars(strip C) \<subseteq> X \<and> (\<forall>a \<in> set(annos C). wt a X))"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    43
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    44
instance ..
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    45
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    46
end
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    47
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    48
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    49
instantiation option :: (wt)wt
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    50
begin
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    51
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    52
definition wt_option where
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    53
"wt opt X = (case opt of None \<Rightarrow> True | Some x \<Rightarrow> wt x X)"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    54
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    55
lemma wt_simps[simp]: "wt None X" "wt (Some x) X = wt x X"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    56
by(simp_all add: wt_option_def)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    57
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    58
instance ..
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    59
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    60
end
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    61
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    62
class SL_top_wt = join + wt +
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    63
fixes top :: "com \<Rightarrow> 'a" ("\<top>\<^bsub>_\<^esub>")
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    64
assumes join_ge1 [simp]: "wt x X \<Longrightarrow> wt y X \<Longrightarrow> x \<sqsubseteq> x \<squnion> y"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    65
and join_ge2 [simp]: "wt x X \<Longrightarrow> wt y X \<Longrightarrow> y \<sqsubseteq> x \<squnion> y"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    66
and join_least[simp]: "x \<sqsubseteq> z \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<squnion> y \<sqsubseteq> z"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    67
and top[simp]: "wt x (vars c) \<Longrightarrow> x \<sqsubseteq> top c"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    68
and wt_top[simp]: "wt (top c) (vars c)"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    69
and wt_join[simp]: "wt x X \<Longrightarrow> wt y X \<Longrightarrow> wt (x \<squnion> y) X"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    70
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    71
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    72
instantiation option :: (SL_top_wt)SL_top_wt
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    73
begin
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    74
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    75
definition top_option where "top c = Some(top c)"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    76
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    77
instance proof
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    78
  case goal1 thus ?case by(cases x, simp, cases y, simp_all)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    79
next
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    80
  case goal2 thus ?case by(cases y, simp, cases x, simp_all)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    81
next
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    82
  case goal3 thus ?case by(cases z, simp, cases y, simp, cases x, simp_all)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    83
next
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    84
  case goal4 thus ?case by(cases x, simp_all add: top_option_def)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    85
next
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    86
  case goal5 thus ?case by(simp add: top_option_def)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    87
next
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    88
  case goal6 thus ?case by(simp add: wt_option_def split: option.splits)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    89
qed
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    90
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    91
end
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    92
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    93
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    94
subsection "Abstract State with Computable Ordering"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    95
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    96
hide_type  st  --"to avoid long names"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    97
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    98
text{* A concrete type of state with computable @{text"\<sqsubseteq>"}: *}
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    99
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   100
datatype 'a st = FunDom "vname \<Rightarrow> 'a" "vname set"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   101
49344
nipkow
parents: 47619
diff changeset
   102
fun "fun" where "fun (FunDom f X) = f"
nipkow
parents: 47619
diff changeset
   103
fun dom where "dom (FunDom f X) = X"
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   104
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   105
definition "show_st S = (\<lambda>x. (x, fun S x)) ` dom S"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   106
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   107
value [code] "show_st (FunDom (\<lambda>x. 1::int) {''a'',''b''})"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   108
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   109
definition "show_acom = map_acom (Option.map show_st)"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   110
definition "show_acom_opt = Option.map show_acom"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   111
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   112
definition "update F x y = FunDom ((fun F)(x:=y)) (dom F)"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   113
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   114
lemma fun_update[simp]: "fun (update S x y) = (fun S)(x:=y)"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   115
by(rule ext)(auto simp: update_def)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   116
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   117
lemma dom_update[simp]: "dom (update S x y) = dom S"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   118
by(simp add: update_def)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   119
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   120
definition "\<gamma>_st \<gamma> F = {f. \<forall>x\<in>dom F. f x \<in> \<gamma>(fun F x)}"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   121
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   122
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   123
instantiation st :: (preord) preord
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   124
begin
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   125
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   126
definition le_st :: "'a st \<Rightarrow> 'a st \<Rightarrow> bool" where
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   127
"F \<sqsubseteq> G = (dom F = dom G \<and> (\<forall>x \<in> dom F. fun F x \<sqsubseteq> fun G x))"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   128
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   129
instance
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   130
proof
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   131
  case goal2 thus ?case by(auto simp: le_st_def)(metis preord_class.le_trans)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   132
qed (auto simp: le_st_def)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   133
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   134
end
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   135
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   136
instantiation st :: (join) join
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   137
begin
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   138
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   139
definition join_st :: "'a st \<Rightarrow> 'a st \<Rightarrow> 'a st" where
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   140
"F \<squnion> G = FunDom (\<lambda>x. fun F x \<squnion> fun G x) (dom F)"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   141
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   142
instance ..
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   143
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   144
end
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   145
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   146
instantiation st :: (type) wt
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   147
begin
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   148
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   149
definition wt_st :: "'a st \<Rightarrow> vname set \<Rightarrow> bool" where
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   150
"wt F X = (dom F = X)"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   151
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   152
instance ..
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   153
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   154
end
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   155
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   156
instantiation st :: (SL_top) SL_top_wt
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   157
begin
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   158
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   159
definition top_st where "top c = FunDom (\<lambda>x. \<top>) (vars c)"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   160
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   161
instance
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   162
proof
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   163
qed (auto simp: le_st_def join_st_def top_st_def wt_st_def)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   164
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   165
end
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   166
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   167
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   168
lemma mono_fun: "F \<sqsubseteq> G \<Longrightarrow> x : dom F \<Longrightarrow> fun F x \<sqsubseteq> fun G x"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   169
by(auto simp: le_st_def)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   170
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   171
lemma mono_update[simp]:
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   172
  "a1 \<sqsubseteq> a2 \<Longrightarrow> S1 \<sqsubseteq> S2 \<Longrightarrow> update S1 x a1 \<sqsubseteq> update S2 x a2"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   173
by(auto simp add: le_st_def update_def)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   174
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   175
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   176
locale Gamma = Val_abs where \<gamma>=\<gamma> for \<gamma> :: "'av::SL_top \<Rightarrow> val set"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   177
begin
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   178
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   179
abbreviation \<gamma>\<^isub>f :: "'av st \<Rightarrow> state set"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   180
where "\<gamma>\<^isub>f == \<gamma>_st \<gamma>"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   181
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   182
abbreviation \<gamma>\<^isub>o :: "'av st option \<Rightarrow> state set"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   183
where "\<gamma>\<^isub>o == \<gamma>_option \<gamma>\<^isub>f"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   184
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   185
abbreviation \<gamma>\<^isub>c :: "'av st option acom \<Rightarrow> state set acom"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   186
where "\<gamma>\<^isub>c == map_acom \<gamma>\<^isub>o"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   187
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   188
lemma gamma_f_Top[simp]: "\<gamma>\<^isub>f (top c) = UNIV"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   189
by(auto simp: top_st_def \<gamma>_st_def)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   190
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   191
lemma gamma_o_Top[simp]: "\<gamma>\<^isub>o (top c) = UNIV"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   192
by (simp add: top_option_def)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   193
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   194
(* FIXME (maybe also le \<rightarrow> sqle?) *)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   195
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   196
lemma mono_gamma_f: "f \<sqsubseteq> g \<Longrightarrow> \<gamma>\<^isub>f f \<subseteq> \<gamma>\<^isub>f g"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   197
apply(simp add:\<gamma>_st_def subset_iff le_st_def split: if_splits)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   198
by (metis mono_gamma subsetD)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   199
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   200
lemma mono_gamma_o:
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   201
  "S1 \<sqsubseteq> S2 \<Longrightarrow> \<gamma>\<^isub>o S1 \<subseteq> \<gamma>\<^isub>o S2"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   202
by(induction S1 S2 rule: le_option.induct)(simp_all add: mono_gamma_f)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   203
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   204
lemma mono_gamma_c: "C1 \<sqsubseteq> C2 \<Longrightarrow> \<gamma>\<^isub>c C1 \<le> \<gamma>\<^isub>c C2"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   205
by (induction C1 C2 rule: le_acom.induct) (simp_all add:mono_gamma_o)
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   206
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   207
lemma in_gamma_option_iff:
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   208
  "x : \<gamma>_option r u \<longleftrightarrow> (\<exists>u'. u = Some u' \<and> x : r u')"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   209
by (cases u) auto
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   210
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   211
end
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   212
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   213
end