src/HOL/MicroJava/BV/Typing_Framework_JVM.thy
author hoelzl
Thu, 12 Nov 2009 17:21:48 +0100
changeset 33639 603320b93668
parent 32693 6c6b1ba5e71e
child 33954 1bc3b688548c
permissions -rw-r--r--
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/BV/JVM.thy
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
     2
    ID:         $Id$
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
     3
    Author:     Tobias Nipkow, Gerwin Klein
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
     4
    Copyright   2000 TUM
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
     5
*)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
     6
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
     7
header {* \isaheader{The Typing Framework for the JVM}\label{sec:JVM} *}
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
     8
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14045
diff changeset
     9
theory Typing_Framework_JVM imports Typing_Framework_err JVMType EffectMono BVSpec begin
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    10
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    11
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    12
constdefs
26450
158b924b5153 avoid amiguity of State.state vs. JVMType.state;
wenzelm
parents: 26289
diff changeset
    13
  exec :: "jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> instr list \<Rightarrow> JVMType.state step_type"
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    14
  "exec G maxs rT et bs == 
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    15
  err_step (size bs) (\<lambda>pc. app (bs!pc) G maxs rT pc et) (\<lambda>pc. eff (bs!pc) G pc et)"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    16
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    17
constdefs
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    18
  opt_states :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> (ty list \<times> ty err list) option set"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    19
  "opt_states G maxs maxr \<equiv> opt (\<Union>{list n (types G) |n. n \<le> maxs} \<times> list maxr (err (types G)))"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    20
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    21
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    22
section {*  Executability of @{term check_bounded} *}
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    23
consts
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    24
  list_all'_rec :: "('a \<Rightarrow> nat \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> bool"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    25
primrec
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    26
  "list_all'_rec P n []     = True"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    27
  "list_all'_rec P n (x#xs) = (P x n \<and> list_all'_rec P (Suc n) xs)"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    28
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    29
constdefs
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    30
  list_all' :: "('a \<Rightarrow> nat \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    31
  "list_all' P xs \<equiv> list_all'_rec P 0 xs"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    32
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    33
lemma list_all'_rec:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    34
  "\<And>n. list_all'_rec P n xs = (\<forall>p < size xs. P (xs!p) (p+n))"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    35
  apply (induct xs)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    36
  apply auto
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    37
  apply (case_tac p)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    38
  apply auto
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    39
  done
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    40
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    41
lemma list_all' [iff]:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    42
  "list_all' P xs = (\<forall>n < size xs. P (xs!n) n)"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    43
  by (unfold list_all'_def) (simp add: list_all'_rec)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    44
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    45
lemma [code]:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    46
  "check_bounded ins et = 
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    47
  (list_all' (\<lambda>i pc. list_all (\<lambda>pc'. pc' < length ins) (succs i pc)) ins \<and> 
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    48
   list_all (\<lambda>e. fst (snd (snd e)) < length ins) et)"
17090
603f23d71ada small mods to code lemmas
nipkow
parents: 16417
diff changeset
    49
  by (simp add: list_all_iff check_bounded_def)
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    50
  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    51
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    52
section {* Connecting JVM and Framework *}
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    53
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    54
lemma check_bounded_is_bounded:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    55
  "check_bounded ins et \<Longrightarrow> bounded (\<lambda>pc. eff (ins!pc) G pc et) (length ins)"  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    56
  by (unfold bounded_def) (blast dest: check_boundedD)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    57
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    58
lemma special_ex_swap_lemma [iff]: 
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    59
  "(? X. (? n. X = A n & P n) & Q X) = (? n. Q(A n) & P n)"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    60
  by blast
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    61
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    62
lemmas [iff del] = not_None_eq
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    63
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    64
theorem exec_pres_type:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    65
  "wf_prog wf_mb S \<Longrightarrow> 
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    66
  pres_type (exec S maxs rT et bs) (size bs) (states S maxs maxr)"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    67
  apply (unfold exec_def JVM_states_unfold)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    68
  apply (rule pres_type_lift)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    69
  apply clarify
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    70
  apply (case_tac s)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    71
   apply simp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    72
   apply (drule effNone)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    73
   apply simp  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    74
  apply (simp add: eff_def xcpt_eff_def norm_eff_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    75
  apply (case_tac "bs!p")
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    76
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    77
  apply (clarsimp simp add: not_Err_eq)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    78
  apply (drule listE_nth_in, assumption)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    79
  apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    80
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    81
  apply (fastsimp simp add: not_None_eq)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    82
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    83
  apply (fastsimp simp add: not_None_eq typeof_empty_is_type)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    84
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    85
  apply clarsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    86
  apply (erule disjE)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    87
   apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    88
  apply clarsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    89
  apply (rule_tac x="1" in exI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    90
  apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    91
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    92
  apply clarsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    93
  apply (erule disjE)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    94
   apply (fastsimp dest: field_fields fields_is_type)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    95
  apply (simp add: match_some_entry split: split_if_asm)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    96
  apply (rule_tac x=1 in exI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    97
  apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    98
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
    99
  apply clarsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   100
  apply (erule disjE)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   101
   apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   102
  apply (simp add: match_some_entry split: split_if_asm)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   103
  apply (rule_tac x=1 in exI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   104
  apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   105
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   106
  apply clarsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   107
  apply (erule disjE)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   108
   apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   109
  apply clarsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   110
  apply (rule_tac x=1 in exI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   111
  apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   112
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   113
  defer 
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   114
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   115
  apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   116
  apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   117
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   118
  apply clarsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   119
  apply (rule_tac x="n'+2" in exI)  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   120
  apply simp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   121
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   122
  apply clarsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   123
  apply (rule_tac x="Suc (Suc (Suc (length ST)))" in exI)  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   124
  apply simp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   125
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   126
  apply clarsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   127
  apply (rule_tac x="Suc (Suc (Suc (Suc (length ST))))" in exI)  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   128
  apply simp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   129
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   130
  apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   131
  apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   132
  apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   133
  apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   134
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   135
  apply clarsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   136
  apply (erule disjE)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   137
   apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   138
  apply clarsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   139
  apply (rule_tac x=1 in exI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   140
  apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   141
  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   142
  apply (erule disjE)
32693
6c6b1ba5e71e tuned proofs
haftmann
parents: 26450
diff changeset
   143
   apply clarsimp
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   144
   apply (drule method_wf_mdecl, assumption+)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   145
   apply (clarsimp simp add: wf_mdecl_def wf_mhead_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   146
   apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   147
  apply clarsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   148
  apply (rule_tac x=1 in exI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   149
  apply fastsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   150
  done
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   151
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   152
lemmas [iff] = not_None_eq
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   153
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   154
lemma sup_state_opt_unfold:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   155
  "sup_state_opt G \<equiv> Opt.le (Product.le (Listn.le (subtype G)) (Listn.le (Err.le (subtype G))))"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   156
  by (simp add: sup_state_opt_def sup_state_def sup_loc_def sup_ty_opt_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   157
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   158
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   159
lemma app_mono:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   160
  "app_mono (sup_state_opt G) (\<lambda>pc. app (bs!pc) G maxs rT pc et) (length bs) (opt_states G maxs maxr)"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   161
  by (unfold app_mono_def lesub_def) (blast intro: EffectMono.app_mono)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   162
  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   163
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   164
lemma list_appendI:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   165
  "\<lbrakk>a \<in> list x A; b \<in> list y A\<rbrakk> \<Longrightarrow> a @ b \<in> list (x+y) A"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   166
  apply (unfold list_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   167
  apply (simp (no_asm))
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   168
  apply blast
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   169
  done
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   170
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   171
lemma list_map [simp]:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   172
  "(map f xs \<in> list (length xs) A) = (f ` set xs \<subseteq> A)"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   173
  apply (unfold list_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   174
  apply simp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   175
  done
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   176
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   177
lemma [iff]:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   178
  "(OK ` A \<subseteq> err B) = (A \<subseteq> B)"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   179
  apply (unfold err_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   180
  apply blast
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   181
  done
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   182
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   183
lemma [intro]:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   184
  "x \<in> A \<Longrightarrow> replicate n x \<in> list n A"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   185
  by (induct n, auto)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   186
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   187
lemma lesubstep_type_simple:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   188
  "a <=[Product.le (op =) r] b \<Longrightarrow> a <=|r| b"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   189
  apply (unfold lesubstep_type_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   190
  apply clarify
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   191
  apply (simp add: set_conv_nth)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   192
  apply clarify
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   193
  apply (drule le_listD, assumption)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   194
  apply (clarsimp simp add: lesub_def Product.le_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   195
  apply (rule exI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   196
  apply (rule conjI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   197
   apply (rule exI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   198
   apply (rule conjI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   199
    apply (rule sym)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   200
    apply assumption
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   201
   apply assumption
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   202
  apply assumption
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   203
  done
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   204
  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   205
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   206
lemma eff_mono:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   207
  "\<lbrakk>p < length bs; s <=_(sup_state_opt G) t; app (bs!p) G maxs rT pc et t\<rbrakk>
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   208
  \<Longrightarrow> eff (bs!p) G p et s <=|sup_state_opt G| eff (bs!p) G p et t"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   209
  apply (unfold eff_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   210
  apply (rule lesubstep_type_simple)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   211
  apply (rule le_list_appendI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   212
   apply (simp add: norm_eff_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   213
   apply (rule le_listI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   214
    apply simp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   215
   apply simp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   216
   apply (simp add: lesub_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   217
   apply (case_tac s)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   218
    apply simp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   219
   apply (simp del: split_paired_All split_paired_Ex)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   220
   apply (elim exE conjE)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   221
   apply simp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   222
   apply (drule eff'_mono, assumption)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   223
   apply assumption
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   224
  apply (simp add: xcpt_eff_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   225
  apply (rule le_listI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   226
    apply simp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   227
  apply simp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   228
  apply (simp add: lesub_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   229
  apply (case_tac s)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   230
   apply simp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   231
  apply simp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   232
  apply (case_tac t)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   233
   apply simp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   234
  apply (clarsimp simp add: sup_state_conv)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   235
  done
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   236
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   237
lemma order_sup_state_opt:
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13601
diff changeset
   238
  "ws_prog G \<Longrightarrow> order (sup_state_opt G)"
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   239
  by (unfold sup_state_opt_unfold) (blast dest: acyclic_subcls1 order_widen)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   240
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   241
theorem exec_mono:
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13601
diff changeset
   242
  "ws_prog G \<Longrightarrow> bounded (exec G maxs rT et bs) (size bs) \<Longrightarrow>
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   243
  mono (JVMType.le G maxs maxr) (exec G maxs rT et bs) (size bs) (states G maxs maxr)"  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   244
  apply (unfold exec_def JVM_le_unfold JVM_states_unfold)  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   245
  apply (rule mono_lift)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   246
     apply (fold sup_state_opt_unfold opt_states_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   247
     apply (erule order_sup_state_opt)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   248
    apply (rule app_mono)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   249
   apply assumption
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   250
  apply clarify
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   251
  apply (rule eff_mono)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   252
  apply assumption+
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   253
  done
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   254
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   255
theorem semilat_JVM_slI:
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13601
diff changeset
   256
  "ws_prog G \<Longrightarrow> semilat (JVMType.sl G maxs maxr)"
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   257
  apply (unfold JVMType.sl_def stk_esl_def reg_sl_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   258
  apply (rule semilat_opt)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   259
  apply (rule err_semilat_Product_esl)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   260
  apply (rule err_semilat_upto_esl)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   261
  apply (rule err_semilat_JType_esl, assumption+)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   262
  apply (rule err_semilat_eslI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   263
  apply (rule Listn_sl)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   264
  apply (rule err_semilat_JType_esl, assumption+)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   265
  done
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   266
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   267
lemma sl_triple_conv:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   268
  "JVMType.sl G maxs maxr == 
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   269
  (states G maxs maxr, JVMType.le G maxs maxr, JVMType.sup G maxs maxr)"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   270
  by (simp (no_asm) add: states_def JVMType.le_def JVMType.sup_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   271
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   272
lemma is_type_pTs:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   273
  "\<lbrakk> wf_prog wf_mb G; (C,S,fs,mdecls) \<in> set G; ((mn,pTs),rT,code) \<in> set mdecls \<rbrakk>
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   274
  \<Longrightarrow> set pTs \<subseteq> types G"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   275
proof 
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   276
  assume "wf_prog wf_mb G" 
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   277
         "(C,S,fs,mdecls) \<in> set G"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   278
         "((mn,pTs),rT,code) \<in> set mdecls"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   279
  hence "wf_mdecl wf_mb G C ((mn,pTs),rT,code)"
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13601
diff changeset
   280
    by (rule wf_prog_wf_mdecl)
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   281
  hence "\<forall>t \<in> set pTs. is_type G t" 
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   282
    by (unfold wf_mdecl_def wf_mhead_def) auto
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   283
  moreover
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   284
  fix t assume "t \<in> set pTs"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   285
  ultimately
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   286
  have "is_type G t" by blast
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   287
  thus "t \<in> types G" ..
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   288
qed
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   289
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   290
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   291
lemma jvm_prog_lift:  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   292
  assumes wf: 
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   293
  "wf_prog (\<lambda>G C bd. P G C bd) G"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   294
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   295
  assumes rule:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   296
  "\<And>wf_mb C mn pTs C rT maxs maxl b et bd.
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   297
   wf_prog wf_mb G \<Longrightarrow>
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   298
   method (G,C) (mn,pTs) = Some (C,rT,maxs,maxl,b,et) \<Longrightarrow>
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   299
   is_class G C \<Longrightarrow>
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   300
   set pTs \<subseteq> types G \<Longrightarrow>
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   301
   bd = ((mn,pTs),rT,maxs,maxl,b,et) \<Longrightarrow>
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   302
   P G C bd \<Longrightarrow>
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   303
   Q G C bd"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   304
 
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   305
  shows 
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   306
  "wf_prog (\<lambda>G C bd. Q G C bd) G"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   307
proof -
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   308
  from wf show ?thesis
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   309
    apply (unfold wf_prog_def wf_cdecl_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   310
    apply clarsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   311
    apply (drule bspec, assumption)
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13601
diff changeset
   312
    apply (unfold wf_cdecl_mdecl_def)
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   313
    apply clarsimp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   314
    apply (drule bspec, assumption)
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13601
diff changeset
   315
    apply (frule methd [OF wf [THEN wf_prog_ws_prog]], assumption+)
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   316
    apply (frule is_type_pTs [OF wf], assumption+)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   317
    apply clarify
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   318
    apply (drule rule [OF wf], assumption+)
26289
9d2c375e242b avoid unclear fact references;
wenzelm
parents: 17090
diff changeset
   319
    apply (rule HOL.refl)
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   320
    apply assumption+
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   321
    done
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   322
qed
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   323
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents:
diff changeset
   324
end