src/HOL/MicroJava/Comp/Index.thy
author eberlm
Wed, 25 May 2016 12:24:00 +0200
changeset 63144 76130b7cc450
parent 63040 eb4ddd18d635
child 69085 9999d7823b8f
permissions -rw-r--r--
NEWS: Permutations of a set and randomised folds
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/Comp/Index.thy
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     2
    Author:     Martin Strecker
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     3
*)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     4
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     5
(* Index of variable in list of parameter names and local variables *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     6
17778
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 16417
diff changeset
     7
theory Index
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 16417
diff changeset
     8
imports AuxLemmas DefsComp
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 16417
diff changeset
     9
begin
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    10
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    11
(*indexing a variable name among all variable declarations in a method body*)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    12
definition index :: "java_mb => vname => nat" where
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    13
 "index ==  \<lambda> (pn,lv,blk,res) v.
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    14
  if v = This
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    15
  then 0 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    16
  else Suc (length (takeWhile (\<lambda> z. z~=v) (pn @ map fst lv)))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    17
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    18
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    19
lemma index_length_pns: "
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    20
  \<lbrakk> i = index (pns,lvars,blk,res) vn;
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    21
  wf_java_mdecl G C ((mn,pTs),rT, (pns,lvars,blk,res));
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    22
  vn \<in> set pns\<rbrakk>
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    23
  \<Longrightarrow> 0 < i \<and> i < Suc (length pns)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    24
  apply (simp add: wf_java_mdecl_def index_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    25
  apply (subgoal_tac "vn \<noteq> This")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    26
   apply (auto intro: length_takeWhile)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    27
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    28
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    29
lemma index_length_lvars: "
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    30
  \<lbrakk> i = index (pns,lvars,blk,res) vn;
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    31
  wf_java_mdecl G C ((mn,pTs),rT, (pns,lvars,blk,res));
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    32
  vn \<in> set (map fst lvars)\<rbrakk>
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    33
  \<Longrightarrow> (length pns) < i \<and> i < Suc((length pns) + (length lvars))"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    34
  apply (simp add: wf_java_mdecl_def index_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    35
  apply (subgoal_tac "vn \<noteq> This")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    36
   apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    37
   apply (subgoal_tac "\<forall> x \<in> set pns. (\<lambda>z. z \<noteq> vn) x")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    38
    apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    39
    apply (subgoal_tac "length (takeWhile (\<lambda>z. z \<noteq> vn) (map fst lvars)) < length (map fst lvars)")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    40
     apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    41
    apply (rule length_takeWhile)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    42
    apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    43
   apply (simp add: map_of_in_set)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    44
   apply (intro strip notI) 
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    45
   apply simp 
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    46
  apply blast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    47
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    48
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    49
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    50
(***  index  ***)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    51
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    52
lemma select_at_index : 
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    53
  "x \<in> set (gjmb_plns (gmb G C S)) \<or> x = This
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    54
  \<Longrightarrow> (the (loc This) # glvs (gmb G C S) loc) ! (index (gmb G C S) x) = the (loc x)"
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    55
  apply (simp only: index_def gjmb_plns_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    56
  apply (case_tac "gmb G C S" rule: prod.exhaust)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    57
  apply (simp add: galldefs del: set_append map_append)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    58
  apply (rename_tac a b)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    59
  apply (case_tac b, simp add: gmb_def gjmb_lvs_def del: set_append map_append)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    60
  apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    61
  apply (simp del: set_append map_append)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    62
  apply (frule length_takeWhile)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    63
  apply (frule_tac f = "(the \<circ> loc)" in nth_map)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    64
  apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    65
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    66
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13737
diff changeset
    67
lemma lift_if: "(f (if b then t else e)) = (if b then (f t) else (f e))"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    68
  by auto
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13737
diff changeset
    69
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    70
lemma update_at_index: "
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    71
  \<lbrakk> distinct (gjmb_plns (gmb G C S));
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    72
  x \<in> set (gjmb_plns (gmb G C S)); x \<noteq> This \<rbrakk> \<Longrightarrow>
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    73
  locvars_xstate G C S (Norm (h, l))[index (gmb G C S) x := val] =
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    74
  locvars_xstate G C S (Norm (h, l(x\<mapsto>val)))"
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    75
  apply (simp only: locvars_xstate_def locvars_locals_def index_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    76
  apply (case_tac "gmb G C S" rule: prod.exhaust, simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    77
  apply (rename_tac a b)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    78
  apply (case_tac b, simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    79
  apply (rule conjI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    80
   apply (simp add: gl_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    81
  apply (simp add: galldefs del: set_append map_append)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    82
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    83
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    84
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    85
(* !!!! incomprehensible: why can't List.takeWhile_append2 be applied the same 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    86
  way in the second case as in the first case ? *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    87
lemma index_of_var: "\<lbrakk> xvar \<notin> set pns; xvar \<notin> set (map fst zs); xvar \<noteq> This \<rbrakk>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    88
  \<Longrightarrow> index (pns, zs @ ((xvar, xval) # xys), blk, res) xvar = Suc (length pns + length zs)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    89
  apply (simp add: index_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    90
  apply (subgoal_tac "(\<And>x. ((x \<in> (set pns)) \<Longrightarrow> ((\<lambda>z. (z \<noteq> xvar))x)))")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    91
   apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    92
   apply (subgoal_tac "(takeWhile (\<lambda>z. z \<noteq> xvar) (map fst zs @ xvar # map fst xys)) = map fst zs @ (takeWhile (\<lambda>z. z \<noteq> xvar) (xvar # map fst xys))")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    93
    apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    94
   apply (rule List.takeWhile_append2)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    95
   apply auto
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
    96
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    97
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    98
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    99
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   100
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 60304
diff changeset
   101
(* The following definition should replace the conditions in WellType.thy / wf_java_mdecl
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   102
*)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   103
definition disjoint_varnames :: "[vname list, (vname \<times> ty) list] \<Rightarrow> bool" where
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   104
  "disjoint_varnames pns lvars \<equiv> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   105
  distinct pns \<and> unique lvars \<and> This \<notin> set pns \<and> This \<notin> set (map fst lvars) \<and> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   106
  (\<forall>pn\<in>set pns. pn \<notin> set (map fst lvars))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   107
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   108
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   109
lemma index_of_var2: "
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   110
  disjoint_varnames pns (lvars_pre @ (vn, ty) # lvars_post)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   111
  \<Longrightarrow> index (pns, lvars_pre @ (vn, ty) # lvars_post, blk, res) vn =
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   112
  Suc (length pns + length lvars_pre)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
   113
  apply (simp add: disjoint_varnames_def index_def unique_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
   114
  apply (subgoal_tac "vn \<noteq> This", simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
   115
   apply (subgoal_tac
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
   116
    "takeWhile (\<lambda>z. z \<noteq> vn) (map fst lvars_pre @ vn # map fst lvars_post) =
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
   117
    map fst lvars_pre @ takeWhile (\<lambda>z. z \<noteq> vn) (vn # map fst lvars_post)")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
   118
    apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
   119
   apply (rule List.takeWhile_append2)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
   120
   apply auto
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
   121
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   122
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   123
lemma wf_java_mdecl_disjoint_varnames: 
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
   124
  "wf_java_mdecl G C (S,rT,(pns,lvars,blk,res))
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   125
  \<Longrightarrow> disjoint_varnames pns lvars"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
   126
  apply (cases S)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
   127
  apply (simp add: wf_java_mdecl_def disjoint_varnames_def  map_of_in_set)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
   128
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   129
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   130
lemma wf_java_mdecl_length_pTs_pns: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   131
  "wf_java_mdecl G C ((mn, pTs), rT, pns, lvars, blk, res)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   132
  \<Longrightarrow> length pTs = length pns"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 55417
diff changeset
   133
  by (simp add: wf_java_mdecl_def)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   134
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   135
end