src/HOL/IMP/Abs_State.thy
author wenzelm
Wed, 10 Jun 2015 18:48:48 +0200
changeset 60417 014d77e5c1ab
parent 55466 786edc984c98
child 61179 16775cad1a5c
permissions -rw-r--r--
prefer direct Assumption.add_assms -- avoid term bindings of Proof_Context.add_assms;
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
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
     7
type_synonym 'a st_rep = "(vname * 'a) list"
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
     8
54930
f2ec28292479 minimized class dependency, updated references
nipkow
parents: 53015
diff changeset
     9
fun fun_rep :: "('a::top) st_rep \<Rightarrow> vname \<Rightarrow> 'a" where
51715
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    10
"fun_rep [] = (\<lambda>x. \<top>)" |
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    11
"fun_rep ((x,a)#ps) = (fun_rep ps) (x := a)"
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    12
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    13
lemma fun_rep_map_of[code]: --"original def is too slow"
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    14
  "fun_rep ps = (%x. case map_of ps x of None \<Rightarrow> \<top> | Some a \<Rightarrow> a)"
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    15
by(induction ps rule: fun_rep.induct) auto
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    16
54930
f2ec28292479 minimized class dependency, updated references
nipkow
parents: 53015
diff changeset
    17
definition eq_st :: "('a::top) st_rep \<Rightarrow> 'a st_rep \<Rightarrow> bool" where
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    18
"eq_st S1 S2 = (fun_rep S1 = fun_rep S2)"
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    19
51720
nipkow
parents: 51715
diff changeset
    20
hide_type st  --"hide previous def to avoid long names"
nipkow
parents: 51715
diff changeset
    21
54930
f2ec28292479 minimized class dependency, updated references
nipkow
parents: 53015
diff changeset
    22
quotient_type 'a st = "('a::top) st_rep" / eq_st
51800
nipkow
parents: 51720
diff changeset
    23
morphisms rep_st St
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    24
by (metis eq_st_def equivpI reflpI sympI transpI)
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    25
54930
f2ec28292479 minimized class dependency, updated references
nipkow
parents: 53015
diff changeset
    26
lift_definition update :: "('a::top) st \<Rightarrow> vname \<Rightarrow> 'a \<Rightarrow> 'a st"
51715
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    27
  is "\<lambda>ps x a. (x,a)#ps"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    28
by(auto simp: eq_st_def)
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    29
54930
f2ec28292479 minimized class dependency, updated references
nipkow
parents: 53015
diff changeset
    30
lift_definition "fun" :: "('a::top) st \<Rightarrow> vname \<Rightarrow> 'a" is fun_rep
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    31
by(simp add: eq_st_def)
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    32
54930
f2ec28292479 minimized class dependency, updated references
nipkow
parents: 53015
diff changeset
    33
definition show_st :: "vname set \<Rightarrow> ('a::top) st \<Rightarrow> (vname * 'a)set" where
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    34
"show_st X S = (\<lambda>x. (x, fun S x)) ` X"
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    35
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 54930
diff changeset
    36
definition "show_acom C = map_acom (map_option (show_st (vars(strip C)))) C"
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 54930
diff changeset
    37
definition "show_acom_opt = map_option show_acom"
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    38
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    39
lemma fun_update[simp]: "fun (update S x y) = (fun S)(x:=y)"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    40
by transfer auto
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    41
54930
f2ec28292479 minimized class dependency, updated references
nipkow
parents: 53015
diff changeset
    42
definition \<gamma>_st :: "(('a::top) \<Rightarrow> 'b set) \<Rightarrow> 'a st \<Rightarrow> (vname \<Rightarrow> 'b) set" where
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    43
"\<gamma>_st \<gamma> F = {f. \<forall>x. f x \<in> \<gamma>(fun F x)}"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    44
54930
f2ec28292479 minimized class dependency, updated references
nipkow
parents: 53015
diff changeset
    45
instantiation st :: (order_top) order
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    46
begin
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    47
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    48
definition less_eq_st_rep :: "'a st_rep \<Rightarrow> 'a st_rep \<Rightarrow> bool" where
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    49
"less_eq_st_rep ps1 ps2 =
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    50
  ((\<forall>x \<in> set(map fst ps1) \<union> set(map fst ps2). fun_rep ps1 x \<le> fun_rep ps2 x))"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    51
51715
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    52
lemma less_eq_st_rep_iff:
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    53
  "less_eq_st_rep r1 r2 = (\<forall>x. fun_rep r1 x \<le> fun_rep r2 x)"
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    54
apply(auto simp: less_eq_st_rep_def fun_rep_map_of split: option.split)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    55
apply (metis Un_iff map_of_eq_None_iff option.distinct(1))
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    56
apply (metis Un_iff map_of_eq_None_iff option.distinct(1))
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    57
done
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    58
51715
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    59
corollary less_eq_st_rep_iff_fun:
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    60
  "less_eq_st_rep r1 r2 = (fun_rep r1 \<le> fun_rep r2)"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    61
by (metis less_eq_st_rep_iff le_fun_def)
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    62
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    63
lift_definition less_eq_st :: "'a st \<Rightarrow> 'a st \<Rightarrow> bool" is less_eq_st_rep
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    64
by(auto simp add: eq_st_def less_eq_st_rep_iff)
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    65
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    66
definition less_st where "F < (G::'a st) = (F \<le> G \<and> \<not> G \<le> F)"
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    67
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    68
instance
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    69
proof
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    70
  case goal1 show ?case by(rule less_st_def)
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    71
next
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    72
  case goal2 show ?case by transfer (auto simp: less_eq_st_rep_def)
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    73
next
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    74
  case goal3 thus ?case
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    75
    by transfer (metis less_eq_st_rep_iff order_trans)
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    76
next
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    77
  case goal4 thus ?case
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    78
    by transfer (metis less_eq_st_rep_iff eq_st_def fun_eq_iff antisym)
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    79
qed
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    80
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    81
end
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
    82
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    83
lemma le_st_iff: "(F \<le> G) = (\<forall>x. fun F x \<le> fun G x)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    84
by transfer (rule less_eq_st_rep_iff)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    85
54930
f2ec28292479 minimized class dependency, updated references
nipkow
parents: 53015
diff changeset
    86
fun map2_st_rep :: "('a::top \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a st_rep \<Rightarrow> 'a st_rep \<Rightarrow> 'a st_rep" where
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    87
"map2_st_rep f [] ps2 = map (%(x,y). (x, f \<top> y)) ps2" |
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    88
"map2_st_rep f ((x,y)#ps1) ps2 =
51715
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    89
  (let y2 = fun_rep ps2 x
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    90
   in (x,f y y2) # map2_st_rep f ps1 ps2)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    91
51715
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    92
lemma fun_rep_map2_rep[simp]: "f \<top> \<top> = \<top> \<Longrightarrow>
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    93
  fun_rep (map2_st_rep f ps1 ps2) = (\<lambda>x. f (fun_rep ps1 x) (fun_rep ps2 x))"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    94
apply(induction f ps1 ps2 rule: map2_st_rep.induct)
51715
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    95
apply(simp add: fun_rep_map_of map_of_map fun_eq_iff split: option.split)
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
    96
apply(fastforce simp: fun_rep_map_of fun_eq_iff split:option.splits)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    97
done
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
    98
51826
054a40461449 canonical names of classes
nipkow
parents: 51800
diff changeset
    99
instantiation st :: (semilattice_sup_top) semilattice_sup_top
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   100
begin
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   101
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   102
lift_definition sup_st :: "'a st \<Rightarrow> 'a st \<Rightarrow> 'a st" is "map2_st_rep (op \<squnion>)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   103
by (simp add: eq_st_def)
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   104
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51826
diff changeset
   105
lift_definition top_st :: "'a st" is "[]" .
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   106
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   107
instance
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   108
proof
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   109
  case goal1 show ?case by transfer (simp add:less_eq_st_rep_iff)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   110
next
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   111
  case goal2 show ?case by transfer (simp add:less_eq_st_rep_iff)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   112
next
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   113
  case goal3 thus ?case by transfer (simp add:less_eq_st_rep_iff)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   114
next
51715
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
   115
  case goal4 show ?case
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
   116
    by transfer (simp add:less_eq_st_rep_iff fun_rep_map_of)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   117
qed
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   118
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   119
end
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   120
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   121
lemma fun_top: "fun \<top> = (\<lambda>x. \<top>)"
51715
17b992f19b51 avoided map_of in def of fun_rep (but still needed for efficient code)
nipkow
parents: 51711
diff changeset
   122
by transfer simp
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   123
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   124
lemma mono_update[simp]:
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   125
  "a1 \<le> a2 \<Longrightarrow> S1 \<le> S2 \<Longrightarrow> update S1 x a1 \<le> update S2 x a2"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   126
by transfer (auto simp add: less_eq_st_rep_def)
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   127
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   128
lemma mono_fun: "S1 \<le> S2 \<Longrightarrow> fun S1 x \<le> fun S2 x"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   129
by transfer (simp add: less_eq_st_rep_iff)
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   130
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52020
diff changeset
   131
locale Gamma_semilattice = Val_semilattice where \<gamma>=\<gamma>
51826
054a40461449 canonical names of classes
nipkow
parents: 51800
diff changeset
   132
  for \<gamma> :: "'av::semilattice_sup_top \<Rightarrow> val set"
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   133
begin
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   134
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   135
abbreviation \<gamma>\<^sub>s :: "'av st \<Rightarrow> state set"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   136
where "\<gamma>\<^sub>s == \<gamma>_st \<gamma>"
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   137
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   138
abbreviation \<gamma>\<^sub>o :: "'av st option \<Rightarrow> state set"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   139
where "\<gamma>\<^sub>o == \<gamma>_option \<gamma>\<^sub>s"
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   140
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   141
abbreviation \<gamma>\<^sub>c :: "'av st option acom \<Rightarrow> state set acom"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   142
where "\<gamma>\<^sub>c == map_acom \<gamma>\<^sub>o"
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   143
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   144
lemma gamma_s_top[simp]: "\<gamma>\<^sub>s \<top> = UNIV"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   145
by(auto simp: \<gamma>_st_def fun_top)
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   146
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   147
lemma gamma_o_Top[simp]: "\<gamma>\<^sub>o \<top> = UNIV"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   148
by (simp add: top_option_def)
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   149
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   150
lemma mono_gamma_s: "f \<le> g \<Longrightarrow> \<gamma>\<^sub>s f \<subseteq> \<gamma>\<^sub>s g"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51698
diff changeset
   151
by(simp add:\<gamma>_st_def le_st_iff subset_iff) (metis mono_gamma subsetD)
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   152
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   153
lemma mono_gamma_o:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   154
  "S1 \<le> S2 \<Longrightarrow> \<gamma>\<^sub>o S1 \<subseteq> \<gamma>\<^sub>o S2"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   155
by(induction S1 S2 rule: less_eq_option.induct)(simp_all add: mono_gamma_s)
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   156
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   157
lemma mono_gamma_c: "C1 \<le> C2 \<Longrightarrow> \<gamma>\<^sub>c C1 \<le> \<gamma>\<^sub>c C2"
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51826
diff changeset
   158
by (simp add: less_eq_acom_def mono_gamma_o size_annos anno_map_acom size_annos_same[of C1 C2])
47619
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   159
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   160
lemma in_gamma_option_iff:
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   161
  "x : \<gamma>_option r u \<longleftrightarrow> (\<exists>u'. u = Some u' \<and> x : r u')"
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   162
by (cases u) auto
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   163
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   164
end
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   165
0d3e95375bb7 forgot to add file
nipkow
parents:
diff changeset
   166
end