src/HOL/MicroJava/Comp/TypeInf.thy
author haftmann
Mon, 01 Mar 2010 13:40:23 +0100
changeset 35416 d8d7d1b785af
parent 17778 93d7e524417a
child 52866 438f578ef1d9
permissions -rw-r--r--
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
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/TypeInf.thy
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     2
    Author:     Martin Strecker
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     3
*)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     4
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     5
(* Exact position in theory hierarchy still to be determined *)
17778
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 16417
diff changeset
     6
theory TypeInf
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 16417
diff changeset
     7
imports "../J/WellType"
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 16417
diff changeset
     8
begin
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     9
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    10
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    11
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    12
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    13
;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    14
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    15
(*** Inversion of typing rules -- to be moved into WellType.thy
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    16
     Also modify the wtpd_expr_\<dots> proofs in CorrComp.thy ***)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    17
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    18
lemma NewC_invers: "E\<turnstile>NewC C::T 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    19
  \<Longrightarrow> T = Class C \<and> is_class (prg E) C"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    20
by (erule ty_expr.cases, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    21
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    22
lemma Cast_invers: "E\<turnstile>Cast D e::T
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13673
diff changeset
    23
  \<Longrightarrow> \<exists> C. T = Class D \<and> E\<turnstile>e::C \<and> is_class (prg E) D \<and> prg E\<turnstile>C\<preceq>? Class D"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    24
by (erule ty_expr.cases, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    25
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    26
lemma Lit_invers: "E\<turnstile>Lit x::T
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    27
  \<Longrightarrow> typeof (\<lambda>v. None) x = Some T"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    28
by (erule ty_expr.cases, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    29
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    30
lemma LAcc_invers: "E\<turnstile>LAcc v::T
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    31
  \<Longrightarrow> localT E v = Some T \<and> is_type (prg E) T"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    32
by (erule ty_expr.cases, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    33
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    34
lemma BinOp_invers: "E\<turnstile>BinOp bop e1 e2::T'
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    35
  \<Longrightarrow> \<exists> T. E\<turnstile>e1::T \<and> E\<turnstile>e2::T \<and> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    36
            (if bop = Eq then T' = PrimT Boolean
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    37
                        else T' = T \<and> T = PrimT Integer)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    38
by (erule ty_expr.cases, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    39
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    40
lemma LAss_invers: "E\<turnstile>v::=e::T'
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    41
  \<Longrightarrow> \<exists> T. v ~= This \<and> E\<turnstile>LAcc v::T \<and> E\<turnstile>e::T' \<and> prg E\<turnstile>T'\<preceq>T"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    42
by (erule ty_expr.cases, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    43
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    44
lemma FAcc_invers: "E\<turnstile>{fd}a..fn::fT
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    45
  \<Longrightarrow> \<exists> C. E\<turnstile>a::Class C \<and> field (prg E,C) fn = Some (fd,fT)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    46
by (erule ty_expr.cases, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    47
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    48
lemma FAss_invers: "E\<turnstile>{fd}a..fn:=v::T' 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    49
\<Longrightarrow> \<exists> T. E\<turnstile>{fd}a..fn::T \<and> E\<turnstile>v ::T' \<and> prg E\<turnstile>T'\<preceq>T"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    50
by (erule ty_expr.cases, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    51
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    52
lemma Call_invers: "E\<turnstile>{C}a..mn({pTs'}ps)::rT
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    53
  \<Longrightarrow> \<exists> pTs md. 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    54
  E\<turnstile>a::Class C \<and> E\<turnstile>ps[::]pTs \<and> max_spec (prg E) C (mn, pTs) = {((md,rT),pTs')}"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    55
by (erule ty_expr.cases, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    56
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    57
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    58
lemma Nil_invers: "E\<turnstile>[] [::] Ts \<Longrightarrow> Ts = []"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    59
by (erule ty_exprs.cases, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    60
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    61
lemma Cons_invers: "E\<turnstile>e#es[::]Ts \<Longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    62
  \<exists> T Ts'. Ts = T#Ts' \<and> E \<turnstile>e::T \<and> E \<turnstile>es[::]Ts'"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    63
by (erule ty_exprs.cases, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    64
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    65
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    66
lemma Expr_invers: "E\<turnstile>Expr e\<surd> \<Longrightarrow> \<exists> T. E\<turnstile>e::T"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    67
by (erule wt_stmt.cases, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    68
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    69
lemma Comp_invers: "E\<turnstile>s1;; s2\<surd> \<Longrightarrow> E\<turnstile>s1\<surd> \<and> E\<turnstile>s2\<surd>"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    70
by (erule wt_stmt.cases, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    71
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    72
lemma Cond_invers: "E\<turnstile>If(e) s1 Else s2\<surd> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    73
  \<Longrightarrow> E\<turnstile>e::PrimT Boolean \<and> E\<turnstile>s1\<surd> \<and> E\<turnstile>s2\<surd>"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    74
by (erule wt_stmt.cases, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    75
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    76
lemma Loop_invers: "E\<turnstile>While(e) s\<surd>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    77
  \<Longrightarrow> E\<turnstile>e::PrimT Boolean \<and> E\<turnstile>s\<surd>"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    78
by (erule wt_stmt.cases, auto)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    79
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    80
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    81
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    82
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    83
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    84
declare split_paired_All [simp del]
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    85
declare split_paired_Ex [simp del]
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    86
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    87
(* Uniqueness of types property *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    88
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    89
lemma uniqueness_of_types: "
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13673
diff changeset
    90
  (\<forall> (E\<Colon>'a prog \<times> (vname \<Rightarrow> ty option)) T1 T2. 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13673
diff changeset
    91
  E\<turnstile>e :: T1 \<longrightarrow> E\<turnstile>e :: T2 \<longrightarrow> T1 = T2) \<and>
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13673
diff changeset
    92
  (\<forall> (E\<Colon>'a prog \<times> (vname \<Rightarrow> ty option)) Ts1 Ts2. 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13673
diff changeset
    93
  E\<turnstile>es [::] Ts1 \<longrightarrow> E\<turnstile>es [::] Ts2 \<longrightarrow> Ts1 = Ts2)"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    94
apply (rule expr.induct)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    95
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    96
(* NewC *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    97
apply (intro strip) 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    98
apply (erule ty_expr.cases) apply simp+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    99
apply (erule ty_expr.cases) apply simp+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   100
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   101
(* Cast *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   102
apply (intro strip) 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   103
apply (erule ty_expr.cases) apply simp+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   104
apply (erule ty_expr.cases) apply simp+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   105
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   106
(* Lit *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   107
apply (intro strip) 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   108
apply (erule ty_expr.cases) apply simp+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   109
apply (erule ty_expr.cases) apply simp+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   110
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   111
(* BinOp *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   112
apply (intro strip)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   113
apply (case_tac binop)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   114
(* Eq *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   115
apply (erule ty_expr.cases) apply simp+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   116
apply (erule ty_expr.cases) apply simp+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   117
(* Add *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   118
apply (erule ty_expr.cases) apply simp+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   119
apply (erule ty_expr.cases) apply simp+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   120
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   121
(* LAcc *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   122
apply (intro strip) 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   123
apply (erule ty_expr.cases) apply simp+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   124
apply (erule ty_expr.cases) apply simp+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   125
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   126
(* LAss *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   127
apply (intro strip) 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   128
apply (erule ty_expr.cases) apply simp+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   129
apply (erule ty_expr.cases) apply simp+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   130
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   131
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   132
(* FAcc *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   133
apply (intro strip)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   134
apply (drule FAcc_invers)+ apply (erule exE)+ 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   135
  apply (subgoal_tac "C = Ca", simp) apply blast
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   136
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   137
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   138
(* FAss *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   139
apply (intro strip)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   140
apply (drule FAss_invers)+ apply (erule exE)+ apply (erule conjE)+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   141
apply (drule FAcc_invers)+ apply (erule exE)+ apply blast 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   142
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   143
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   144
(* Call *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   145
apply (intro strip)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   146
apply (drule Call_invers)+ apply (erule exE)+ apply (erule conjE)+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   147
apply (subgoal_tac "pTs = pTsa", simp) apply blast
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   148
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   149
(* expression lists *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   150
apply (intro strip)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   151
apply (erule ty_exprs.cases)+ apply simp+
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   152
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   153
apply (intro strip)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   154
apply (erule ty_exprs.cases, simp)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   155
apply (erule ty_exprs.cases, simp)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   156
apply (subgoal_tac "e = ea", simp) apply simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   157
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   158
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   159
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   160
lemma uniqueness_of_types_expr [rule_format (no_asm)]: "
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   161
  (\<forall> E T1 T2. E\<turnstile>e :: T1 \<longrightarrow> E\<turnstile>e :: T2 \<longrightarrow> T1 = T2)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   162
by (rule uniqueness_of_types [THEN conjunct1])
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   163
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   164
lemma uniqueness_of_types_exprs [rule_format (no_asm)]: "
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   165
  (\<forall> E Ts1 Ts2. E\<turnstile>es [::] Ts1 \<longrightarrow> E\<turnstile>es [::] Ts2 \<longrightarrow> Ts1 = Ts2)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   166
by (rule uniqueness_of_types [THEN conjunct2])
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   167
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   168
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   169
  
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   170
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 17778
diff changeset
   171
definition inferred_tp :: "[java_mb env, expr] \<Rightarrow> ty" where
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   172
  "inferred_tp E e == (SOME T. E\<turnstile>e :: T)"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 17778
diff changeset
   173
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 17778
diff changeset
   174
definition inferred_tps :: "[java_mb env, expr list] \<Rightarrow> ty list" where
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   175
  "inferred_tps E es == (SOME Ts. E\<turnstile>es [::] Ts)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   176
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   177
(* get inferred type(s) for well-typed term *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   178
lemma inferred_tp_wt: "E\<turnstile>e :: T \<Longrightarrow> (inferred_tp E e) = T"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   179
by (auto simp: inferred_tp_def intro: uniqueness_of_types_expr)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   180
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   181
lemma inferred_tps_wt: "E\<turnstile>es [::] Ts \<Longrightarrow> (inferred_tps E es) = Ts"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   182
by (auto simp: inferred_tps_def intro: uniqueness_of_types_exprs)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   183
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   184
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   185
end
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   186