| author | blanchet | 
| Sat, 08 Sep 2012 21:30:31 +0200 | |
| changeset 49222 | cbe8c859817c | 
| parent 46742 | 125e49d04cf6 | 
| child 55417 | 01fbfb60c33e | 
| permissions | -rw-r--r-- | 
| 13673 | 1  | 
(* Title: HOL/MicroJava/Comp/Index.thy  | 
2  | 
Author: Martin Strecker  | 
|
3  | 
*)  | 
|
4  | 
||
5  | 
(* Index of variable in list of parameter names and local variables *)  | 
|
6  | 
||
| 17778 | 7  | 
theory Index  | 
8  | 
imports AuxLemmas DefsComp  | 
|
9  | 
begin  | 
|
| 13673 | 10  | 
|
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 | 13  | 
"index == \<lambda> (pn,lv,blk,res) v.  | 
14  | 
if v = This  | 
|
15  | 
then 0  | 
|
16  | 
else Suc (length (takeWhile (\<lambda> z. z~=v) (pn @ map fst lv)))"  | 
|
17  | 
||
18  | 
||
19  | 
lemma index_length_pns: "  | 
|
20  | 
\<lbrakk> i = index (pns,lvars,blk,res) vn;  | 
|
21  | 
wf_java_mdecl G C ((mn,pTs),rT, (pns,lvars,blk,res));  | 
|
22  | 
vn \<in> set pns\<rbrakk>  | 
|
23  | 
\<Longrightarrow> 0 < i \<and> i < Suc (length pns)"  | 
|
24  | 
apply (simp add: wf_java_mdecl_def index_def)  | 
|
25  | 
apply (subgoal_tac "vn \<noteq> This")  | 
|
26  | 
apply (auto intro: length_takeWhile)  | 
|
27  | 
done  | 
|
28  | 
||
29  | 
lemma index_length_lvars: "  | 
|
30  | 
\<lbrakk> i = index (pns,lvars,blk,res) vn;  | 
|
31  | 
wf_java_mdecl G C ((mn,pTs),rT, (pns,lvars,blk,res));  | 
|
32  | 
vn \<in> set (map fst lvars)\<rbrakk>  | 
|
33  | 
\<Longrightarrow> (length pns) < i \<and> i < Suc((length pns) + (length lvars))"  | 
|
34  | 
apply (simp add: wf_java_mdecl_def index_def)  | 
|
35  | 
apply (subgoal_tac "vn \<noteq> This")  | 
|
36  | 
apply simp  | 
|
37  | 
apply (subgoal_tac "\<forall> x \<in> set pns. (\<lambda>z. z \<noteq> vn) x")  | 
|
| 46226 | 38  | 
apply simp  | 
| 13673 | 39  | 
apply (subgoal_tac "length (takeWhile (\<lambda>z. z \<noteq> vn) (map fst lvars)) < length (map fst lvars)")  | 
40  | 
apply simp  | 
|
41  | 
apply (rule length_takeWhile)  | 
|
42  | 
apply simp  | 
|
43  | 
apply (simp add: map_of_in_set)  | 
|
44  | 
apply (intro strip notI) apply simp apply blast  | 
|
45  | 
done  | 
|
46  | 
||
47  | 
||
48  | 
(*** index ***)  | 
|
49  | 
||
50  | 
lemma select_at_index :  | 
|
51  | 
"x \<in> set (gjmb_plns (gmb G C S)) \<or> x = This  | 
|
52  | 
\<Longrightarrow> (the (loc This) # glvs (gmb G C S) loc) ! (index (gmb G C S) x) =  | 
|
53  | 
the (loc x)"  | 
|
54  | 
apply (simp only: index_def gjmb_plns_def)  | 
|
| 27117 | 55  | 
apply (case_tac "gmb G C S" rule: prod.exhaust)  | 
| 13673 | 56  | 
apply (simp add: galldefs del: set_append map_append)  | 
57  | 
apply (case_tac b, simp add: gmb_def gjmb_lvs_def del: set_append map_append)  | 
|
58  | 
apply (intro strip)  | 
|
59  | 
apply (simp del: set_append map_append)  | 
|
60  | 
apply (frule length_takeWhile)  | 
|
61  | 
apply (frule_tac f = "(the \<circ> loc)" in nth_map)  | 
|
62  | 
apply simp  | 
|
63  | 
done  | 
|
64  | 
||
| 14045 | 65  | 
lemma lift_if: "(f (if b then t else e)) = (if b then (f t) else (f e))"  | 
66  | 
apply auto  | 
|
67  | 
done  | 
|
68  | 
||
| 13673 | 69  | 
lemma update_at_index: "  | 
70  | 
\<lbrakk> distinct (gjmb_plns (gmb G C S));  | 
|
71  | 
x \<in> set (gjmb_plns (gmb G C S)); x \<noteq> This \<rbrakk> \<Longrightarrow>  | 
|
72  | 
locvars_xstate G C S (Norm (h, l))[index (gmb G C S) x := val] =  | 
|
73  | 
locvars_xstate G C S (Norm (h, l(x\<mapsto>val)))"  | 
|
74  | 
apply (simp only: locvars_xstate_def locvars_locals_def index_def)  | 
|
| 27117 | 75  | 
apply (case_tac "gmb G C S" rule: prod.exhaust, simp)  | 
| 13673 | 76  | 
apply (case_tac b, simp)  | 
77  | 
apply (rule conjI)  | 
|
78  | 
apply (simp add: gl_def)  | 
|
79  | 
apply (simp add: galldefs del: set_append map_append)  | 
|
80  | 
done  | 
|
81  | 
||
82  | 
||
83  | 
(* !!!! incomprehensible: why can't List.takeWhile_append2 be applied the same  | 
|
84  | 
way in the second case as in the first case ? *)  | 
|
85  | 
lemma index_of_var: "\<lbrakk> xvar \<notin> set pns; xvar \<notin> set (map fst zs); xvar \<noteq> This \<rbrakk>  | 
|
86  | 
\<Longrightarrow> index (pns, zs @ ((xvar, xval) # xys), blk, res) xvar = Suc (length pns + length zs)"  | 
|
87  | 
apply (simp add: index_def)  | 
|
88  | 
apply (subgoal_tac "(\<And>x. ((x \<in> (set pns)) \<Longrightarrow> ((\<lambda>z. (z \<noteq> xvar))x)))")  | 
|
| 46226 | 89  | 
apply simp  | 
| 13673 | 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))")  | 
91  | 
apply simp  | 
|
92  | 
apply (rule List.takeWhile_append2)  | 
|
93  | 
apply auto  | 
|
94  | 
done  | 
|
95  | 
||
96  | 
||
97  | 
||
98  | 
||
99  | 
(* The following def should replace the conditions in WellType.thy / wf_java_mdecl  | 
|
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 | 102  | 
(* This corresponds to the original def in wf_java_mdecl:  | 
103  | 
"disjoint_varnames pns lvars \<equiv>  | 
|
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 | 106  | 
*)  | 
107  | 
||
108  | 
"disjoint_varnames pns lvars \<equiv>  | 
|
109  | 
distinct pns \<and> unique lvars \<and> This \<notin> set pns \<and> This \<notin> set (map fst lvars) \<and>  | 
|
110  | 
(\<forall>pn\<in>set pns. pn \<notin> set (map fst lvars))"  | 
|
111  | 
||
112  | 
||
113  | 
lemma index_of_var2: "  | 
|
114  | 
disjoint_varnames pns (lvars_pre @ (vn, ty) # lvars_post)  | 
|
115  | 
\<Longrightarrow> index (pns, lvars_pre @ (vn, ty) # lvars_post, blk, res) vn =  | 
|
116  | 
Suc (length pns + length lvars_pre)"  | 
|
| 46226 | 117  | 
apply (simp add: disjoint_varnames_def index_def unique_def)  | 
| 13673 | 118  | 
apply (subgoal_tac "vn \<noteq> This", simp)  | 
119  | 
apply (subgoal_tac  | 
|
120  | 
"takeWhile (\<lambda>z. z \<noteq> vn) (map fst lvars_pre @ vn # map fst lvars_post) =  | 
|
121  | 
map fst lvars_pre @ takeWhile (\<lambda>z. z \<noteq> vn) (vn # map fst lvars_post)")  | 
|
122  | 
apply simp  | 
|
123  | 
apply (rule List.takeWhile_append2)  | 
|
124  | 
apply auto  | 
|
125  | 
done  | 
|
126  | 
||
127  | 
lemma wf_java_mdecl_disjoint_varnames:  | 
|
128  | 
"wf_java_mdecl G C (S,rT,(pns,lvars,blk,res))  | 
|
129  | 
\<Longrightarrow> disjoint_varnames pns lvars"  | 
|
| 46742 | 130  | 
apply (cases S)  | 
| 13673 | 131  | 
apply (simp add: wf_java_mdecl_def disjoint_varnames_def map_of_in_set)  | 
132  | 
done  | 
|
133  | 
||
134  | 
lemma wf_java_mdecl_length_pTs_pns:  | 
|
135  | 
"wf_java_mdecl G C ((mn, pTs), rT, pns, lvars, blk, res)  | 
|
136  | 
\<Longrightarrow> length pTs = length pns"  | 
|
137  | 
by (simp add: wf_java_mdecl_def)  | 
|
138  | 
||
139  | 
end  |