src/HOL/MicroJava/Comp/LemmasComp.thy
author paulson
Wed, 16 Jun 2004 14:56:39 +0200
changeset 14952 47455995693d
parent 14045 a34d89ce6097
child 14981 e73f8140af78
permissions -rw-r--r--
removal of x-symbol syntax <Sigma> for dependent products
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/LemmasComp.thy
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     2
    ID:         $Id$
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     3
    Author:     Martin Strecker
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     4
    Copyright   GPL 2002
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     5
*)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     6
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     7
(* Lemmas for compiler correctness proof *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     8
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     9
theory LemmasComp = TranslComp:
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    10
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    11
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    12
declare split_paired_All [simp del]
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    13
declare split_paired_Ex [simp del]
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    14
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    15
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    16
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    17
(* misc lemmas *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    18
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    19
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    20
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    21
lemma split_pairs: "(\<lambda>(a,b). (F a b)) (ab) = F (fst ab) (snd ab)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    22
proof -
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    23
  have "(\<lambda>(a,b). (F a b)) (fst ab,snd ab) = F (fst ab) (snd ab)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    24
    by (simp add: split_def)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    25
  then show ?thesis by simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    26
qed
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    27
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    28
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    29
lemma c_hupd_conv: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    30
  "c_hupd h' (xo, (h,l)) = (xo, (if xo = None then h' else h),l)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    31
by (simp add: c_hupd_def)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    32
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    33
lemma gl_c_hupd [simp]: "(gl (c_hupd h xs)) = (gl xs)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    34
by (simp add: gl_def c_hupd_def split_pairs)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    35
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    36
lemma c_hupd_xcpt_invariant [simp]: "gx (c_hupd h' (xo, st)) = xo"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    37
by (case_tac st, simp only: c_hupd_conv gx_conv)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    38
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    39
(* not added to simpset because of interference with c_hupd_conv *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    40
lemma c_hupd_hp_invariant: "gh (c_hupd hp (None, st)) = hp"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    41
by (case_tac st, simp add: c_hupd_conv gh_def)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    42
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    43
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    44
lemma unique_map_fst [rule_format]: "(\<forall> x \<in> set xs. (fst x = fst (f x) )) \<longrightarrow>
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    45
  unique (map f xs) = unique xs"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    46
proof (induct xs)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    47
  case Nil show ?case by simp
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    48
next
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    49
  case (Cons a list)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    50
  show ?case
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    51
  proof
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    52
    assume fst_eq: "\<forall>x\<in>set (a # list). fst x = fst (f x)"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    53
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    54
    have fst_set: "(fst a \<in> fst ` set list) = (fst (f a) \<in> fst ` f ` set list)" 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    55
    proof 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    56
      assume as: "fst a \<in> fst ` set list" 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    57
      then obtain x where fst_a_x: "x\<in>set list \<and> fst a = fst x" 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    58
	by (auto simp add: image_iff)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    59
      then have "fst (f a) = fst (f x)" by (simp add: fst_eq)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    60
      with as show "(fst (f a) \<in> fst ` f ` set list)" by (simp add: fst_a_x)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    61
    next
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    62
      assume as: "fst (f a) \<in> fst ` f ` set list"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    63
      then obtain x where fst_a_x: "x\<in>set list \<and> fst (f a) = fst (f x)"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    64
	by (auto simp add: image_iff)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    65
      then have "fst a = fst x" by (simp add: fst_eq)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    66
      with as show "fst a \<in> fst ` set list" by (simp add: fst_a_x)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    67
    qed
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    68
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    69
    with fst_eq Cons 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    70
    show "unique (map f (a # list)) = unique (a # list)"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    71
      by (simp add: unique_def fst_set)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    72
  qed
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    73
qed
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    74
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    75
lemma comp_unique: "unique (comp G) = unique G"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    76
apply (simp add: comp_def)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    77
apply (rule unique_map_fst)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    78
apply (simp add: compClass_def split_beta)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    79
done
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    80
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    81
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    82
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    83
(* invariance of properties under compilation *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    84
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    85
lemma comp_class_imp:
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    86
  "(class G C = Some(D, fs, ms)) \<Longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    87
  (class (comp G) C = Some(D, fs, map (compMethod G C) ms))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    88
apply (simp add: class_def comp_def compClass_def)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    89
apply (rule HOL.trans)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    90
apply (rule map_of_map2)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    91
apply auto
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    92
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    93
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    94
lemma comp_class_None: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    95
"(class G C = None) = (class (comp G) C = None)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    96
apply (simp add: class_def comp_def compClass_def)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    97
apply (simp add: map_of_in_set)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    98
apply (simp add: image_compose [THEN sym] o_def split_beta del: image_compose)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    99
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   100
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   101
lemma comp_is_class: "is_class (comp G) C = is_class G C"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   102
by (case_tac "class G C", auto simp: is_class_def comp_class_None dest: comp_class_imp)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   103
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   104
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   105
lemma comp_is_type: "is_type (comp G) T = is_type G T"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   106
  by ((cases T),simp,(induct G)) ((simp),(simp only: comp_is_class),(simp add: comp_is_class),(simp only: comp_is_class))
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   107
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   108
lemma comp_classname: "is_class G C 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   109
  \<Longrightarrow> fst (the (class G C)) = fst (the (class (comp G) C))"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   110
by (case_tac "class G C", auto simp: is_class_def dest: comp_class_imp)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   111
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   112
lemma comp_subcls1: "subcls1 (comp G) = subcls1 G"
14952
47455995693d removal of x-symbol syntax <Sigma> for dependent products
paulson
parents: 14045
diff changeset
   113
by (auto simp add: subcls1_def2 comp_classname comp_is_class)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   114
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   115
lemma comp_widen: "((ty1,ty2) \<in> widen (comp G)) = ((ty1,ty2) \<in> widen G)"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   116
  apply rule
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   117
  apply (cases "(ty1,ty2)" "comp G" rule: widen.cases) 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   118
  apply (simp_all add: comp_subcls1 widen.null)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   119
  apply (cases "(ty1,ty2)" G rule: widen.cases) 
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   120
  apply (simp_all add: comp_subcls1 widen.null)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   121
  done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   122
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   123
lemma comp_cast: "((ty1,ty2) \<in> cast (comp G)) = ((ty1,ty2) \<in> cast G)"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   124
  apply rule
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   125
  apply (cases "(ty1,ty2)" "comp G" rule: cast.cases) 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   126
  apply (simp_all add: comp_subcls1 cast.widen cast.subcls)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   127
  apply (rule cast.widen) apply (simp add: comp_widen)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   128
  apply (cases "(ty1,ty2)" G rule: cast.cases)
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   129
  apply (simp_all add: comp_subcls1 cast.widen cast.subcls)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   130
  apply (rule cast.widen) apply (simp add: comp_widen)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   131
  done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   132
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   133
lemma comp_cast_ok: "cast_ok (comp G) = cast_ok G"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   134
  by (simp add: expand_fun_eq cast_ok_def comp_widen)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   135
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   136
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   137
lemma compClass_fst [simp]: "(fst (compClass G C)) = (fst C)"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   138
by (simp add: compClass_def split_beta)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   139
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   140
lemma compClass_fst_snd [simp]: "(fst (snd (compClass G C))) = (fst (snd C))"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   141
by (simp add: compClass_def split_beta)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   142
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   143
lemma compClass_fst_snd_snd [simp]: "(fst (snd (snd (compClass G C)))) = (fst (snd (snd C)))"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   144
by (simp add: compClass_def split_beta)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   145
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   146
lemma comp_wf_fdecl [simp]: "wf_fdecl (comp G) fd = wf_fdecl G fd"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   147
by (case_tac fd, simp add: wf_fdecl_def comp_is_type)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   148
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   149
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   150
lemma compClass_forall [simp]: "
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   151
  (\<forall>x\<in>set (snd (snd (snd (compClass G C)))). P (fst x) (fst (snd x))) =
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   152
  (\<forall>x\<in>set (snd (snd (snd C))). P (fst x) (fst (snd x)))"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   153
by (simp add: compClass_def compMethod_def split_beta)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   154
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   155
lemma comp_wf_mhead: "wf_mhead (comp G) S rT =  wf_mhead G S rT"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   156
by (simp add: wf_mhead_def split_beta comp_is_type)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   157
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   158
lemma comp_ws_cdecl: "
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   159
  ws_cdecl (TranslComp.comp G) (compClass G C) = ws_cdecl G C"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   160
apply (simp add: ws_cdecl_def split_beta comp_is_class comp_subcls1)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   161
apply (simp (no_asm_simp) add: comp_wf_mhead)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   162
apply (simp add: compClass_def compMethod_def split_beta unique_map_fst)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   163
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   164
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   165
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   166
lemma comp_wf_syscls: "wf_syscls (comp G) = wf_syscls G"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   167
apply (simp add: wf_syscls_def comp_def compClass_def split_beta)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   168
apply (simp only: image_compose [THEN sym])
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   169
apply (subgoal_tac "(Fun.comp fst (\<lambda>(C, cno::cname, fdls::fdecl list, jmdls).
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   170
  (C, cno, fdls, map (compMethod G C) jmdls))) = fst")
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   171
apply (simp del: image_compose)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   172
apply (simp add: expand_fun_eq split_beta)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   173
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   174
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   175
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   176
lemma comp_ws_prog: "ws_prog (comp G) = ws_prog G"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   177
apply (rule sym)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   178
apply (simp add: ws_prog_def comp_ws_cdecl comp_unique)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   179
apply (simp add: comp_wf_syscls)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   180
apply (auto simp add: comp_ws_cdecl [THEN sym] TranslComp.comp_def)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   181
done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   182
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   183
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   184
lemma comp_class_rec: " wf ((subcls1 G)^-1) \<Longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   185
class_rec (comp G) C t f = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   186
  class_rec G C t (\<lambda> C' fs' ms' r'. f C' fs' (map (compMethod G C') ms') r')"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   187
apply (rule_tac a = C in  wf_induct) apply assumption
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   188
apply (subgoal_tac "wf ((subcls1 (comp G))\<inverse>)")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   189
apply (subgoal_tac "(class G x = None) \<or> (\<exists> D fs ms. (class G x = Some (D, fs, ms)))")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   190
apply (erule disjE)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   191
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   192
  (* case class G x = None *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   193
apply (simp (no_asm_simp) add: class_rec_def comp_subcls1 wfrec cut_apply)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   194
apply (simp add: comp_class_None)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   195
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   196
  (* case \<exists> D fs ms. (class G x = Some (D, fs, ms)) *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   197
apply (erule exE)+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   198
apply (frule comp_class_imp)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   199
apply (frule_tac G="comp G" and C=x and t=t and f=f in class_rec_lemma)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   200
  apply assumption
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   201
apply (frule_tac G=G and C=x and t=t 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   202
   and f="(\<lambda>C' fs' ms'. f C' fs' (map (compMethod G C') ms'))" in class_rec_lemma)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   203
  apply assumption
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   204
apply (simp only:)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   205
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   206
apply (case_tac "x = Object")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   207
  apply simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   208
  apply (frule subcls1I, assumption)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   209
    apply (drule_tac x=D in spec, drule mp, simp)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   210
    apply simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   211
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   212
  (* subgoals *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   213
apply (case_tac "class G x")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   214
apply auto
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   215
apply (simp add: comp_subcls1)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   216
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   217
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   218
lemma comp_fields: "wf ((subcls1 G)^-1) \<Longrightarrow> 
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   219
  fields (comp G,C) = fields (G,C)" 
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   220
by (simp add: fields_def comp_class_rec)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   221
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   222
lemma comp_field: "wf ((subcls1 G)^-1) \<Longrightarrow> 
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   223
  field (comp G,C) = field (G,C)" 
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   224
by (simp add: field_def comp_fields)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   225
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   226
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   227
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   228
lemma class_rec_relation [rule_format (no_asm)]: "\<lbrakk>  ws_prog G;
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   229
  \<forall> fs ms. R (f1 Object fs ms t1) (f2 Object fs ms t2);
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   230
   \<forall> C fs ms r1 r2. (R r1 r2) \<longrightarrow> (R (f1 C fs ms r1) (f2 C fs ms r2)) \<rbrakk>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   231
  \<Longrightarrow> ((class G C) \<noteq> None) \<longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   232
  R (class_rec G C t1 f1) (class_rec G C t2 f2)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   233
apply (frule wf_subcls1) (* establish wf ((subcls1 G)^-1) *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   234
apply (rule_tac a = C in  wf_induct) apply assumption
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   235
apply (intro strip)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   236
apply (subgoal_tac "(\<exists>D rT mb. class G x = Some (D, rT, mb))")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   237
  apply (erule exE)+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   238
apply (frule_tac C=x and t=t1 and f=f1 in class_rec_lemma)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   239
  apply assumption
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   240
apply (frule_tac C=x and t=t2 and f=f2 in class_rec_lemma)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   241
  apply assumption
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   242
apply (simp only:)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   243
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   244
apply (case_tac "x = Object")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   245
  apply simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   246
  apply (frule subcls1I, assumption)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   247
    apply (drule_tac x=D in spec, drule mp, simp)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   248
    apply simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   249
    apply (subgoal_tac "(\<exists>D' rT' mb'. class G D = Some (D', rT', mb'))")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   250
    apply blast
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   251
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   252
  (* subgoals *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   253
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   254
apply (frule class_wf_struct) apply assumption
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   255
apply (simp add: ws_cdecl_def is_class_def)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   256
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   257
apply (simp add: subcls1_def2 is_class_def)
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   258
apply auto
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   259
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   260
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   261
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   262
syntax
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   263
  mtd_mb :: "cname \<times> ty \<times> 'c \<Rightarrow> 'c"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   264
translations
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   265
  "mtd_mb" => "Fun.comp snd snd"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   266
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   267
lemma map_of_map_fst: "\<lbrakk> inj f;
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   268
  \<forall>x\<in>set xs. fst (f x) = fst x; \<forall>x\<in>set xs. fst (g x) = fst x \<rbrakk>
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   269
  \<Longrightarrow>  map_of (map g xs) k 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   270
  = option_map (\<lambda> e. (snd (g ((inv f) (k, e))))) (map_of (map f xs) k)"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   271
apply (induct xs)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   272
apply simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   273
apply (simp del: split_paired_All)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   274
apply (case_tac "k = fst a")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   275
apply (simp del: split_paired_All)
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   276
apply (subgoal_tac "(inv f (fst a, snd (f a))) = a", simp)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   277
apply (subgoal_tac "(fst a, snd (f a)) = f a", simp)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   278
apply (erule conjE)+
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   279
apply (drule_tac s ="fst (f a)" and t="fst a" in sym)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   280
apply simp
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   281
apply (simp add: surjective_pairing)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   282
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   283
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   284
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   285
lemma comp_method [rule_format (no_asm)]: "\<lbrakk> ws_prog G; is_class G C\<rbrakk> \<Longrightarrow> 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   286
  ((method (comp G, C) S) = 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   287
  option_map (\<lambda> (D,rT,b).  (D, rT, mtd_mb (compMethod G D (S, rT, b))))
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   288
             (method (G, C) S))"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   289
apply (simp add: method_def)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   290
apply (frule wf_subcls1)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   291
apply (simp add: comp_class_rec)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   292
apply (simp add: map_compose [THEN sym] split_iter split_compose del: map_compose)
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   293
apply (rule_tac R="%x y. ((x S) = (option_map (\<lambda>(D, rT, b). 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   294
  (D, rT, snd (snd (compMethod G D (S, rT, b))))) (y S)))" 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   295
  in class_rec_relation)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   296
apply assumption
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   297
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   298
apply (intro strip)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   299
apply simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   300
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   301
apply (rule trans)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   302
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   303
apply (rule_tac f="(\<lambda>(s, m). (s, Object, m))" and
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   304
  g="(Fun.comp (\<lambda>(s, m). (s, Object, m)) (compMethod G Object))" 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   305
  in map_of_map_fst)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   306
defer				(* inj \<dots> *)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   307
apply (simp add: inj_on_def split_beta) 
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   308
apply (simp add: split_beta compMethod_def)
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   309
apply (simp add: map_of_map [THEN sym])
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   310
apply (simp add: split_beta)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   311
apply (simp add: map_compose [THEN sym] Fun.comp_def split_beta)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   312
apply (subgoal_tac "(\<lambda>x\<Colon>(vname list \<times> fdecl list \<times> stmt \<times> expr) mdecl.
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   313
                    (fst x, Object,
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   314
                     snd (compMethod G Object
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   315
                           (inv (\<lambda>(s\<Colon>sig, m\<Colon>ty \<times> vname list \<times> fdecl list \<times> stmt \<times> expr).
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   316
                                    (s, Object, m))
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   317
                             (S, Object, snd x)))))
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   318
  = (\<lambda>x. (fst x, Object, fst (snd x),
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   319
                        snd (snd (compMethod G Object (S, snd x)))))")
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   320
apply (simp only:)
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   321
apply (simp add: expand_fun_eq)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   322
apply (intro strip)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   323
apply (subgoal_tac "(inv (\<lambda>(s, m). (s, Object, m)) (S, Object, snd x)) = (S, snd x)")
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   324
apply (simp only:)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   325
apply (simp add: compMethod_def split_beta)
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   326
apply (rule inv_f_eq) 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   327
defer
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   328
defer
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   329
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   330
apply (intro strip)
14025
d9b155757dc8 *** empty log message ***
nipkow
parents: 13673
diff changeset
   331
apply (simp add: map_add_Some_iff map_of_map del: split_paired_Ex)
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   332
apply (simp add: map_add_def)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   333
apply (subgoal_tac "inj (\<lambda>(s, m). (s, Ca, m))")
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   334
apply (drule_tac g="(Fun.comp (\<lambda>(s, m). (s, Ca, m)) (compMethod G Ca))" and xs=ms
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   335
  and k=S in map_of_map_fst) 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   336
apply (simp add: split_beta) 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   337
apply (simp add: compMethod_def split_beta)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   338
apply (case_tac "(map_of (map (\<lambda>(s, m). (s, Ca, m)) ms) S)")
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   339
apply simp
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   340
apply simp apply (simp add: split_beta map_of_map) apply (erule exE) apply (erule conjE)+
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   341
apply (drule_tac t=a in sym)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   342
apply (subgoal_tac "(inv (\<lambda>(s, m). (s, Ca, m)) (S, a)) = (S, snd a)")
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   343
apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   344
apply (subgoal_tac "\<forall>x\<in>set ms. fst ((Fun.comp (\<lambda>(s, m). (s, Ca, m)) (compMethod G Ca)) x) = fst x")
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   345
   prefer 2 apply (simp (no_asm_simp) add: compMethod_def split_beta)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   346
apply (simp add: map_of_map2)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   347
apply (simp (no_asm_simp) add: compMethod_def split_beta)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   348
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   349
  -- "remaining subgoals"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   350
apply (auto intro: inv_f_eq simp add: inj_on_def is_class_def)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   351
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   352
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   353
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   354
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   355
lemma comp_wf_mrT: "\<lbrakk> ws_prog G; is_class G D\<rbrakk> \<Longrightarrow> 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   356
  wf_mrT (TranslComp.comp G) (C, D, fs, map (compMethod G a) ms) =
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   357
  wf_mrT G (C, D, fs, ms)"
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   358
apply (simp add: wf_mrT_def compMethod_def split_beta)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   359
apply (simp add: comp_widen)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   360
apply (rule iffI)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   361
apply (intro strip)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   362
apply simp
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   363
apply (drule bspec) apply assumption
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   364
apply (drule_tac x=D' in spec) apply (drule_tac x=rT' in spec) apply (drule mp)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   365
prefer 2 apply assumption
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   366
apply (simp add: comp_method [of G D])
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   367
apply (erule exE)+
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   368
apply (subgoal_tac "z =(fst z, fst (snd z), snd (snd z))")
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   369
apply (rule exI)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   370
apply (simp)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   371
apply (simp add: split_paired_all)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   372
apply (intro strip)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   373
apply (simp add: comp_method)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   374
apply auto
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   375
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   376
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   377
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   378
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   379
  (* DIVERSE OTHER LEMMAS *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   380
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   381
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   382
lemma max_spec_preserves_length:
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   383
  "max_spec G C (mn, pTs) = {((md,rT),pTs')} 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   384
  \<Longrightarrow> length pTs = length pTs'"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   385
apply (frule max_spec2mheads)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   386
apply (erule exE)+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   387
apply (simp add: list_all2_def)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   388
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   389
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   390
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   391
lemma ty_exprs_length [simp]: "(E\<turnstile>es[::]Ts \<longrightarrow> length es = length Ts)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   392
apply (subgoal_tac "(E\<turnstile>e::T \<longrightarrow> True) \<and> (E\<turnstile>es[::]Ts \<longrightarrow> length es = length Ts) \<and> (E\<turnstile>s\<surd> \<longrightarrow> True)")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   393
apply blast
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   394
apply (rule ty_expr_ty_exprs_wt_stmt.induct)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   395
apply auto
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   396
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   397
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   398
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   399
lemma max_spec_preserves_method_rT [simp]:
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   400
  "max_spec G C (mn, pTs) = {((md,rT),pTs')}
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   401
  \<Longrightarrow> method_rT (the (method (G, C) (mn, pTs'))) = rT"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   402
apply (frule max_spec2mheads)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   403
apply (erule exE)+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   404
apply (simp add: method_rT_def)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   405
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   406
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   407
  (**********************************************************************************)
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   408
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   409
declare compClass_fst [simp del]
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   410
declare compClass_fst_snd [simp del]
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   411
declare compClass_fst_snd_snd [simp del]
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   412
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   413
declare split_paired_All [simp add]
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   414
declare split_paired_Ex [simp add]
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   415
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   416
end