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