src/HOL/MicroJava/Comp/Index.thy
author wenzelm
Sat, 07 Apr 2012 16:41:59 +0200
changeset 47389 e8552cba702d
parent 46742 125e49d04cf6
child 55417 01fbfb60c33e
permissions -rw-r--r--
explicit checks stable_finished_theory/stable_command allow parallel asynchronous command transactions; tuned;
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;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    21
  wf_java_mdecl G C ((mn,pTs),rT, (pns,lvars,blk,res)); 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    22
  vn \<in> set pns\<rbrakk> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    23
  \<Longrightarrow> 0 < i \<and> i < Suc (length pns)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    24
apply (simp add: wf_java_mdecl_def index_def)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    25
apply (subgoal_tac "vn \<noteq> This")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    26
apply (auto intro: length_takeWhile)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    27
done
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;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    31
  wf_java_mdecl G C ((mn,pTs),rT, (pns,lvars,blk,res)); 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    32
  vn \<in> set (map fst lvars)\<rbrakk> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    33
  \<Longrightarrow> (length pns) < i \<and> i < Suc((length pns) + (length lvars))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    34
apply (simp add: wf_java_mdecl_def index_def)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    35
apply (subgoal_tac "vn \<noteq> This")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    36
apply simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    37
apply (subgoal_tac "\<forall> x \<in> set pns. (\<lambda>z. z \<noteq> vn) x")
46226
e88e980ed735 tuned proofs;
wenzelm
parents: 35416
diff changeset
    38
apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    39
apply (subgoal_tac "length (takeWhile (\<lambda>z. z \<noteq> vn) (map fst lvars)) < length (map fst lvars)")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    40
apply simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    41
apply (rule length_takeWhile)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    42
apply simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    43
apply (simp add: map_of_in_set)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    44
apply (intro strip notI) apply simp apply blast
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    45
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    46
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    47
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    48
(***  index  ***)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    49
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    50
lemma select_at_index : 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    51
  "x \<in> set (gjmb_plns (gmb G C S)) \<or> x = This 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    52
  \<Longrightarrow> (the (loc This) # glvs (gmb G C S) loc) ! (index (gmb G C S) x) = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    53
     the (loc x)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    54
apply (simp only: index_def gjmb_plns_def)
27117
97e9dae57284 case_split_tac (works without context);
wenzelm
parents: 17778
diff changeset
    55
apply (case_tac "gmb G C S" rule: prod.exhaust)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    56
apply (simp add: galldefs del: set_append map_append)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    57
apply (case_tac b, simp add: gmb_def gjmb_lvs_def del: set_append map_append)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    58
apply (intro strip)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    59
apply (simp del: set_append map_append)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    60
apply (frule length_takeWhile)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    61
apply (frule_tac f = "(the \<circ> loc)" in nth_map)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    62
apply simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    63
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    64
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13737
diff changeset
    65
lemma lift_if: "(f (if b then t else e)) = (if b then (f t) else (f e))"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13737
diff changeset
    66
apply auto
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13737
diff changeset
    67
done
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13737
diff changeset
    68
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    69
lemma update_at_index: "
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    70
  \<lbrakk> distinct (gjmb_plns (gmb G C S)); 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    71
  x \<in> set (gjmb_plns (gmb G C S)); x \<noteq> This \<rbrakk> \<Longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    72
  locvars_xstate G C S (Norm (h, l))[index (gmb G C S) x := val] =
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    73
          locvars_xstate G C S (Norm (h, l(x\<mapsto>val)))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    74
apply (simp only: locvars_xstate_def locvars_locals_def index_def)
27117
97e9dae57284 case_split_tac (works without context);
wenzelm
parents: 17778
diff changeset
    75
apply (case_tac "gmb G C S" rule: prod.exhaust, simp)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    76
apply (case_tac b, simp)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    77
apply (rule conjI)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    78
apply (simp add: gl_def)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    79
apply (simp add: galldefs del: set_append map_append)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    80
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    81
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    82
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    83
(* !!!! incomprehensible: why can't List.takeWhile_append2 be applied the same 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    84
  way in the second case as in the first case ? *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    85
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
    86
  \<Longrightarrow> index (pns, zs @ ((xvar, xval) # xys), blk, res) xvar = Suc (length pns + length zs)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    87
apply (simp add: index_def)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    88
apply (subgoal_tac "(\<And>x. ((x \<in> (set pns)) \<Longrightarrow> ((\<lambda>z. (z \<noteq> xvar))x)))")
46226
e88e980ed735 tuned proofs;
wenzelm
parents: 35416
diff changeset
    89
apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    90
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))")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    91
apply simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    92
apply (rule List.takeWhile_append2)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    93
apply auto
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    94
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    95
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    96
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
(* The following def should replace the conditions in WellType.thy / wf_java_mdecl
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   100
*)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   101
definition disjoint_varnames :: "[vname list, (vname \<times> ty) list] \<Rightarrow> bool" where
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   102
(* This corresponds to the original def in wf_java_mdecl:
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   103
  "disjoint_varnames pns lvars \<equiv> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   104
  nodups pns \<and> unique lvars \<and> This \<notin> set pns \<and> This \<notin> set (map fst lvars) \<and> 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27117
diff changeset
   105
        (\<forall>pn\<in>set pns. map_of lvars pn = None)"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   106
*)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   107
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   108
  "disjoint_varnames pns lvars \<equiv> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   109
  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
   110
  (\<forall>pn\<in>set pns. pn \<notin> set (map fst lvars))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   111
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   112
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   113
lemma index_of_var2: "
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   114
  disjoint_varnames pns (lvars_pre @ (vn, ty) # lvars_post)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   115
  \<Longrightarrow> index (pns, lvars_pre @ (vn, ty) # lvars_post, blk, res) vn =
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   116
  Suc (length pns + length lvars_pre)"
46226
e88e980ed735 tuned proofs;
wenzelm
parents: 35416
diff changeset
   117
apply (simp add: disjoint_varnames_def index_def unique_def)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   118
apply (subgoal_tac "vn \<noteq> This", simp)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   119
apply (subgoal_tac
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   120
  "takeWhile (\<lambda>z. z \<noteq> vn) (map fst lvars_pre @ vn # map fst lvars_post) =
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   121
  map fst lvars_pre @ takeWhile (\<lambda>z. z \<noteq> vn) (vn # map fst lvars_post)")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   122
apply simp 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   123
apply (rule List.takeWhile_append2)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   124
apply auto
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   125
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   126
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   127
lemma wf_java_mdecl_disjoint_varnames: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   128
  "wf_java_mdecl G C (S,rT,(pns,lvars,blk,res)) 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   129
  \<Longrightarrow> disjoint_varnames pns lvars"
46742
125e49d04cf6 tuned proofs;
wenzelm
parents: 46226
diff changeset
   130
apply (cases S)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   131
apply (simp add: wf_java_mdecl_def disjoint_varnames_def  map_of_in_set)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   132
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   133
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   134
lemma wf_java_mdecl_length_pTs_pns: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   135
  "wf_java_mdecl G C ((mn, pTs), rT, pns, lvars, blk, res)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   136
  \<Longrightarrow> length pTs = length pns"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   137
by (simp add: wf_java_mdecl_def)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   138
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   139
end