thys/FMap.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Tue, 29 Apr 2014 15:26:48 +0100
changeset 19 087d82632852
child 20 e04123f4bacc
permissions -rw-r--r--
added FMap theory and adapted tm-theory
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     1
theory FMap
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     2
  imports Main "~~/src/HOL/Quotient_Examples/FSet" "~~/src/HOL/Library/DAList"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     3
begin
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
subsubsection {* The type of finite maps *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     6
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     7
typedef ('a, 'b) fmap  (infixr "f\<rightharpoonup>" 1) = "{x :: 'a \<rightharpoonup> 'b. finite (dom x) }"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     8
  proof show "empty \<in> {x. finite (dom x)}" by simp qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     9
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    10
setup_lifting type_definition_fmap
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    11
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    12
lift_definition fdom :: "'key f\<rightharpoonup> 'value \<Rightarrow> 'key set" is "dom" ..
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    13
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    14
lift_definition fran :: "'key f\<rightharpoonup> 'value \<Rightarrow> 'value set" is "ran" ..
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    15
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    16
lift_definition lookup :: "'key f\<rightharpoonup> 'value \<Rightarrow> 'key \<Rightarrow> 'value option" is "(\<lambda> x. x)" ..
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    17
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    18
abbreviation the_lookup (infix "f!" 55)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
  where "m f! x \<equiv> the (lookup m x)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    20
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
lift_definition fempty :: "'key f\<rightharpoonup> 'value" ("f\<emptyset>") is Map.empty by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
lemma fempty_fdom[simp]: "fdom f\<emptyset> = {}"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
  by (transfer, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
lemma fdomIff: "(a : fdom m) = (lookup m a \<noteq> None)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
 by (transfer, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    29
lemma lookup_not_fdom: "x \<notin> fdom m \<Longrightarrow> lookup m x = None"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    30
  by (auto iff:fdomIff)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    31
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
lemma finite_range:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    33
  assumes "finite (dom m)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    34
  shows "finite (ran m)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    35
  apply (rule finite_subset[OF _ finite_imageI[OF assms, of "\<lambda> x . the (m x)"]])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    36
  by (auto simp add: ran_def dom_def image_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    37
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    38
lemma finite_fdom[simp]: "finite (fdom m)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    39
  by transfer
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    40
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    41
lemma finite_fran[simp]: "finite (fran m)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    42
  by (transfer, rule finite_range)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    43
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    44
lemma fmap_eqI[intro]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    45
  assumes "fdom a = fdom b"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    46
  and "\<And> x. x \<in> fdom a \<Longrightarrow> a f! x = b f! x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    47
  shows "a = b"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    48
using assms
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    49
proof(transfer)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    50
  fix a b :: "'a \<rightharpoonup> 'b"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    51
  assume d: "dom a = dom b"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    52
  assume eq: "\<And> x. x \<in> dom a \<Longrightarrow> the (a x) = the (b x)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    53
  show "a = b"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    54
  proof
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    55
    fix x
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    56
    show "a x = b x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    57
    proof(cases "a x")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    58
    case None
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    59
      hence "x \<notin> dom a" by (simp add: dom_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    60
      hence "x \<notin> dom b" using d by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    61
      hence " b x = None"  by (simp add: dom_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    62
      thus ?thesis using None by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    63
    next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    64
    case (Some y)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    65
      hence d': "x \<in> dom ( a)" by (simp add: dom_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    66
      hence "the ( a x) = the ( b x)" using eq by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    67
      moreover
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    68
      have "x \<in> dom ( b)" using Some d' d by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    69
      then obtain y' where " b x = Some y'" by (auto simp add: dom_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    70
      ultimately
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    71
      show " a x =  b x" using Some by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    72
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    73
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    74
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    75
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    76
subsubsection {* Updates *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    77
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    78
lift_definition
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    79
  fmap_upd :: "'key f\<rightharpoonup> 'value \<Rightarrow> 'key \<Rightarrow> 'value \<Rightarrow> 'key f\<rightharpoonup> 'value" ("_'(_ f\<mapsto> _')" [900,900]900)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    80
  is "\<lambda> m x v. m( x \<mapsto> v)"  by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    81
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    82
lemma fmap_upd_fdom[simp]: "fdom (h (x f\<mapsto> v)) = insert x (fdom h)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    83
  by (transfer, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    84
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    85
lemma the_lookup_fmap_upd[simp]: "lookup (h (x f\<mapsto> v)) x = Some v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    86
  by (transfer, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    87
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    88
lemma the_lookup_fmap_upd_other[simp]: "x' \<noteq> x \<Longrightarrow> lookup (h (x f\<mapsto> v)) x' = lookup h x'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
  by (transfer, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    90
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    91
lemma fmap_upd_overwrite[simp]: "f (x f\<mapsto> y) (x f\<mapsto> z) = f (x f\<mapsto> z)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    92
  by (transfer, auto) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    93
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    94
lemma fmap_upd_twist: "a \<noteq> c \<Longrightarrow> (m(a f\<mapsto> b))(c f\<mapsto> d) = (m(c f\<mapsto> d))(a f\<mapsto> b)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    95
  apply (rule fmap_eqI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    96
  apply auto[1]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    97
  apply (case_tac "x = a", auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    98
  apply (case_tac "x = c", auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    99
  done
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   100
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   101
subsubsection {* Restriction *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   102
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   103
lift_definition fmap_restr :: "'a set \<Rightarrow> 'a f\<rightharpoonup> 'b \<Rightarrow> 'a f\<rightharpoonup> 'b"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   104
  is "\<lambda> S m. (if finite S then (restrict_map m S) else empty)" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   105
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   106
lemma lookup_fmap_restr[simp]: "finite S \<Longrightarrow> x \<in> S \<Longrightarrow> lookup (fmap_restr S m) x = lookup m x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   107
  by (transfer, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   108
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   109
lemma fdom_fmap_restr[simp]: "finite S \<Longrightarrow> fdom (fmap_restr S m) = fdom m \<inter> S"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   110
  by (transfer, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   111
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   112
lemma fmap_restr_fmap_restr[simp]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   113
 "finite d1 \<Longrightarrow> finite d2 \<Longrightarrow> fmap_restr d1 (fmap_restr d2 x) = fmap_restr (d1 \<inter> d2) x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   114
 by (transfer, auto simp add: restrict_map_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   115
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   116
lemma fmap_restr_fmap_restr_subset:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   117
 "finite d2 \<Longrightarrow> d1 \<subseteq> d2 \<Longrightarrow> fmap_restr d1 (fmap_restr d2 x) = fmap_restr d1 x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   118
 by (metis Int_absorb2 finite_subset fmap_restr_fmap_restr)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   119
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   120
lemma fmap_restr_useless: "finite S \<Longrightarrow> fdom m \<subseteq> S \<Longrightarrow> fmap_restr S m = m"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   121
  by (rule fmap_eqI, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   122
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   123
lemma fmap_restr_not_finite:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   124
  "\<not> finite S \<Longrightarrow> fmap_restr S \<rho> = f\<emptyset>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   125
  by (transfer, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   126
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   127
lemma fmap_restr_fmap_upd: "x \<in> S \<Longrightarrow> finite S \<Longrightarrow> fmap_restr S (m1(x f\<mapsto> v)) = (fmap_restr S m1)(x f\<mapsto> v)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   128
  apply (rule fmap_eqI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   129
  apply auto[1]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   130
  apply (case_tac "xa = x")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   131
  apply auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   132
  done
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   133
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   134
subsubsection {* Deleting *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   135
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   136
lift_definition fmap_delete :: "'a \<Rightarrow> 'a f\<rightharpoonup> 'b \<Rightarrow> 'a f\<rightharpoonup> 'b"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   137
  is "\<lambda> x m. m(x := None)" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   138
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   139
lemma fdom_fmap_delete[simp]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   140
  "fdom (fmap_delete x m) = fdom m - {x}"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   141
  by (transfer, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   142
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   143
lemma fmap_delete_fmap_upd[simp]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   144
  "fmap_delete x (m(x f\<mapsto> v)) = fmap_delete x m"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   145
  by (transfer, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   146
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   147
lemma fmap_delete_noop:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   148
  "x \<notin> fdom m \<Longrightarrow> fmap_delete x m = m"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   149
  by (transfer, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   150
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   151
lemma fmap_upd_fmap_delete[simp]: "x \<in> fdom \<Gamma> \<Longrightarrow> (fmap_delete x \<Gamma>)(x f\<mapsto> \<Gamma> f! x) = \<Gamma>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   152
  by (transfer, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   153
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   154
lemma fran_fmap_upd[simp]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   155
  "fran (m(x f\<mapsto> v)) = insert v (fran (fmap_delete x m))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   156
by (transfer, auto simp add: ran_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   157
 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   158
subsubsection {* Addition (merging) of finite maps *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   159
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   160
lift_definition fmap_add :: "'a f\<rightharpoonup> 'b \<Rightarrow> 'a f\<rightharpoonup> 'b \<Rightarrow> 'a f\<rightharpoonup> 'b" (infixl "f++" 100) 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   161
  is "map_add" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   162
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   163
lemma fdom_fmap_add[simp]: "fdom (m1 f++ m2) = fdom m1 \<union> fdom m2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   164
  by (transfer, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   165
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   166
lemma lookup_fmap_add1[simp]: "x \<in> fdom m2 \<Longrightarrow> lookup (m1 f++ m2) x = lookup m2 x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   167
  by (transfer, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   168
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   169
lemma lookup_fmap_add2[simp]:  "x \<notin> fdom m2 \<Longrightarrow> lookup (m1 f++ m2) x = lookup m1 x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   170
  apply transfer
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   171
  by (metis map_add_dom_app_simps(3))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   172
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   173
lemma [simp]: "\<rho> f++ f\<emptyset> = \<rho>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   174
  by (transfer, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   175
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   176
lemma fmap_add_overwrite: "fdom m1 \<subseteq> fdom m2 \<Longrightarrow> m1 f++ m2 = m2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   177
  apply transfer
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   178
  apply rule
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   179
  apply (case_tac "x \<in> dom m2")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   180
  apply (auto simp add: map_add_dom_app_simps(1))
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   181
  done
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   182
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   183
lemma fmap_add_upd_swap: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   184
  "x \<notin> fdom \<rho>' \<Longrightarrow> \<rho>(x f\<mapsto> z) f++ \<rho>' = (\<rho> f++ \<rho>')(x f\<mapsto> z)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   185
  apply transfer
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   186
  by (metis map_add_upd_left)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   187
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   188
lemma fmap_add_upd: 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   189
  "\<rho> f++ (\<rho>'(x f\<mapsto> z)) = (\<rho> f++ \<rho>')(x f\<mapsto> z)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   190
  apply transfer
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   191
  by (metis map_add_upd)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   192
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   193
lemma fmap_restr_add: "fmap_restr S (m1 f++ m2) = fmap_restr S m1 f++ fmap_restr S m2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   194
  apply (cases "finite S")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   195
  apply (rule fmap_eqI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   196
  apply auto[1]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   197
  apply (case_tac "x \<in> fdom m2")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   198
  apply auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   199
  apply (simp add: fmap_restr_not_finite)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   200
  done
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   201
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   202
subsubsection {* Conversion from associative lists *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   203
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   204
lift_definition fmap_of :: "('a \<times> 'b) list \<Rightarrow> 'a f\<rightharpoonup> 'b"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   205
  is map_of by (rule finite_dom_map_of)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   206
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   207
lemma fmap_of_Nil[simp]: "fmap_of [] = f\<emptyset>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   208
 by (transfer, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   209
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   210
lemma fmap_of_Cons[simp]: "fmap_of (p # l) = (fmap_of l)(fst p f\<mapsto> snd p)" 
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   211
  by (transfer, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   212
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   213
lemma fmap_of_append[simp]: "fmap_of (l1 @ l2) = fmap_of l2 f++ fmap_of l1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   214
  by (transfer, simp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   215
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   216
lemma lookup_fmap_of[simp]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   217
  "lookup (fmap_of m) x = map_of m x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   218
  by (transfer, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   219
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   220
lemma fmap_delete_fmap_of[simp]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   221
  "fmap_delete x (fmap_of m) = fmap_of (AList.delete x m)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   222
  by (transfer, metis delete_conv')
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   223
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   224
subsubsection {* Less-than-relation for extending finite maps *}
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   225
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   226
instantiation fmap :: (type,type) order
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   227
begin
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   228
  definition "\<rho> \<le> \<rho>' = ((fdom \<rho> \<subseteq> fdom \<rho>') \<and> (\<forall>x \<in> fdom \<rho>. lookup \<rho> x = lookup \<rho>' x))"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   229
  definition "(\<rho>::'a f\<rightharpoonup> 'b) < \<rho>' = (\<rho> \<noteq> \<rho>' \<and> \<rho> \<le> \<rho>')"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   230
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   231
  lemma fmap_less_eqI[intro]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   232
    assumes assm1: "fdom \<rho> \<subseteq> fdom \<rho>'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   233
      and assm2:  "\<And> x. \<lbrakk> x \<in> fdom \<rho> ; x \<in> fdom \<rho>'  \<rbrakk> \<Longrightarrow> \<rho> f! x = \<rho>' f! x "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   234
     shows "\<rho> \<le> \<rho>'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   235
   unfolding less_eq_fmap_def
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   236
   apply rule
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   237
   apply fact
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   238
   apply rule+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   239
   apply (frule subsetD[OF `_ \<subseteq> _`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   240
   apply (frule  assm2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   241
   apply (auto iff: fdomIff)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   242
   done
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   243
  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   244
  lemma fmap_less_eqD:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   245
    assumes "\<rho> \<le> \<rho>'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   246
    assumes "x \<in> fdom \<rho>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   247
    shows "lookup \<rho> x = lookup \<rho>' x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   248
    using assms
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   249
    unfolding less_eq_fmap_def by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   250
  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   251
  lemma fmap_antisym: assumes  "(x:: 'a f\<rightharpoonup> 'b) \<le> y" and "y \<le> x" shows "x = y "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   252
  proof(rule fmap_eqI[rule_format])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   253
      show "fdom x = fdom y" using `x \<le> y` and `y \<le> x` unfolding less_eq_fmap_def by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   254
      
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   255
      fix v
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   256
      assume "v \<in> fdom x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   257
      hence "v \<in> fdom y" using `fdom _ = _` by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   258
  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   259
      thus "x f! v = y f! v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   260
        using `x \<le> y` `v \<in> fdom x` unfolding less_eq_fmap_def by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   261
    qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   262
  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   263
  lemma fmap_trans: assumes  "(x:: 'a f\<rightharpoonup> 'b) \<le> y" and "y \<le> z" shows "x \<le> z"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   264
  proof
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   265
    show "fdom x \<subseteq> fdom z" using `x \<le> y` and `y \<le> z` unfolding less_eq_fmap_def
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   266
      by -(rule order_trans [of _ "fdom y"], auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   267
    
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   268
    fix v
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   269
    assume "v \<in> fdom x" and "v \<in> fdom z"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   270
    hence "v \<in> fdom y" using `x \<le> y`  unfolding less_eq_fmap_def by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   271
    hence "lookup y v = lookup x v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   272
      using `x \<le> y` `v \<in> fdom x` unfolding less_eq_fmap_def by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   273
    moreover
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   274
    have "lookup y v = lookup z v"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   275
      by (rule fmap_less_eqD[OF `y \<le> z`  `v \<in> fdom y`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   276
    ultimately
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   277
    show "x f! v = z f! v" by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   278
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   279
  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   280
  instance
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   281
    apply default
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   282
    using fmap_antisym apply (auto simp add: less_fmap_def)[1]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   283
    apply (auto simp add: less_eq_fmap_def)[1]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   284
    using fmap_trans apply assumption
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   285
    using fmap_antisym apply assumption
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   286
    done
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   287
end
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   288
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   289
lemma fmap_less_fdom:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   290
  "\<rho>1 \<le> \<rho>2 \<Longrightarrow> fdom \<rho>1 \<subseteq> fdom \<rho>2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   291
  by (metis less_eq_fmap_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   292
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   293
lemma fmap_less_restrict:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   294
  "\<rho>1 \<le> \<rho>2 \<longleftrightarrow> \<rho>1 = fmap_restr (fdom \<rho>1) \<rho>2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   295
  unfolding less_eq_fmap_def
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   296
  apply transfer
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   297
  apply (auto simp add:restrict_map_def split:option.split_asm)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   298
  by (metis option.simps(3))+
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   299
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   300
lemma fmap_restr_less:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   301
  "fmap_restr S \<rho> \<le> \<rho>"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   302
  unfolding less_eq_fmap_def
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   303
  by (transfer, auto)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   304
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   305
lemma fmap_upd_less[simp, intro]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   306
  "\<rho>1 \<le> \<rho>2 \<Longrightarrow> v1 = v2 \<Longrightarrow> \<rho>1(x f\<mapsto> v1) \<le> \<rho>2(x f\<mapsto> v2)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   307
  apply (rule fmap_less_eqI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   308
  apply (auto dest: fmap_less_fdom)[1]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   309
  apply (case_tac "xa = x")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   310
  apply (auto dest: fmap_less_eqD)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   311
  done
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   312
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   313
lemma fmap_update_less[simp, intro]:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   314
  "\<rho>1 \<le> \<rho>1' \<Longrightarrow> \<rho>2 \<le> \<rho>2' \<Longrightarrow>  (fdom \<rho>2' - fdom \<rho>2) \<inter> fdom \<rho>1 = {} \<Longrightarrow> \<rho>1 f++ \<rho>2 \<le> \<rho>1' f++ \<rho>2'"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   315
  apply (rule fmap_less_eqI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   316
  apply (auto dest: fmap_less_fdom)[1]
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   317
  apply (case_tac "x \<in> fdom \<rho>2")
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   318
  apply (auto dest: fmap_less_eqD fmap_less_fdom)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   319
  apply (metis fmap_less_eqD fmap_less_fdom lookup_fmap_add1 set_mp)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   320
  by (metis Diff_iff Diff_triv fmap_less_eqD lookup_fmap_add2)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   321
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   322
lemma fmap_restr_le:
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   323
  assumes "\<rho>1 \<le> \<rho>2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   324
  assumes "S1 \<subseteq> S2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   325
  assumes [simp]:"finite S2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   326
  shows "fmap_restr S1 \<rho>1 \<le> fmap_restr S2 \<rho>2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   327
proof-
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   328
  have [simp]: "finite S1"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   329
    by (rule finite_subset[OF `S1 \<subseteq> S2` `finite S2`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   330
  show ?thesis
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   331
  proof (rule fmap_less_eqI)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   332
    have "fdom \<rho>1 \<subseteq> fdom \<rho>2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   333
      by (metis assms(1) less_eq_fmap_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   334
    thus "fdom (fmap_restr S1 \<rho>1) \<subseteq> fdom (fmap_restr S2 \<rho>2)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   335
      using `S1 \<subseteq> S2`
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   336
      by auto
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   337
  next
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   338
    fix x
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   339
    assume "x \<in> fdom (fmap_restr S1 \<rho>1) "
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   340
    hence [simp]:"x \<in> fdom \<rho>1" and [simp]:"x \<in> S1" and [simp]: "x \<in> S2"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   341
      by (auto intro: set_mp[OF `S1 \<subseteq> S2`])
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   342
    have "\<rho>1 f! x = \<rho>2 f! x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   343
      by (metis `x \<in> fdom \<rho>1` assms(1) less_eq_fmap_def)
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   344
    thus "(fmap_restr S1 \<rho>1) f! x = (fmap_restr S2 \<rho>2) f! x"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   345
      by simp
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   346
  qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   347
qed
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   348
  
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   349
definition fmapdom_to_list :: "('a :: linorder) f\<rightharpoonup> 'b \<Rightarrow> 'a list"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   350
where
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   351
  "fmapdom_to_list f = (THE xs. set xs = fdom f \<and> sorted xs \<and> distinct xs)"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   352
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   353
definition fmap_to_alist :: "('a :: linorder) f\<rightharpoonup> 'b \<Rightarrow> ('a \<times> 'b) list"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   354
where
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   355
  "fmap_to_alist f = [(x, f f! x). x \<leftarrow> fmapdom_to_list f]"
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   356
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   357
087d82632852 added FMap theory and adapted tm-theory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   358
end