src/HOL/IMP/Types.thy
author paulson
Wed, 17 Apr 2024 22:07:21 +0100
changeset 80130 8262d4f63b58
parent 69597 ff784d5a5bfb
child 80914 d97fdabd9e2b
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 58310
diff changeset
     1
section "A Typed Language"
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
     2
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
     3
theory Types imports Star Complex_Main begin
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
     4
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67406
diff changeset
     5
text \<open>We build on \<^theory>\<open>Complex_Main\<close> instead of \<^theory>\<open>Main\<close> to access
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
     6
the real numbers.\<close>
52405
3dd63180cdbf Explain to beginners why Complex_Main
kleing
parents: 52046
diff changeset
     7
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
     8
subsection "Arithmetic Expressions"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
     9
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    10
datatype val = Iv int | Rv real
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    11
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45200
diff changeset
    12
type_synonym vname = string
e87feee00a4c renamed name -> vname
nipkow
parents: 45200
diff changeset
    13
type_synonym state = "vname \<Rightarrow> val"
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    14
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
    15
text_raw\<open>\snip{aexptDef}{0}{2}{%\<close>
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    16
datatype aexp =  Ic int | Rc real | V vname | Plus aexp aexp
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
    17
text_raw\<open>}%endsnip\<close>
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    18
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    19
inductive taval :: "aexp \<Rightarrow> state \<Rightarrow> val \<Rightarrow> bool" where
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    20
"taval (Ic i) s (Iv i)" |
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    21
"taval (Rc r) s (Rv r)" |
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    22
"taval (V x) s (s x)" |
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    23
"taval a1 s (Iv i1) \<Longrightarrow> taval a2 s (Iv i2)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    24
 \<Longrightarrow> taval (Plus a1 a2) s (Iv(i1+i2))" |
44020
376c1e7b320c fixed wrong isubs in IMP/Types
kleing
parents: 43150
diff changeset
    25
"taval a1 s (Rv r1) \<Longrightarrow> taval a2 s (Rv r2)
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    26
 \<Longrightarrow> taval (Plus a1 a2) s (Rv(r1+r2))"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    27
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    28
inductive_cases [elim!]:
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    29
  "taval (Ic i) s v"  "taval (Rc i) s v"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    30
  "taval (V x) s v"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    31
  "taval (Plus a1 a2) s v"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    32
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    33
subsection "Boolean Expressions"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    34
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    35
datatype bexp = Bc bool | Not bexp | And bexp bexp | Less aexp aexp
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    36
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    37
inductive tbval :: "bexp \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> bool" where
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    38
"tbval (Bc v) s v" |
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    39
"tbval b s bv \<Longrightarrow> tbval (Not b) s (\<not> bv)" |
44020
376c1e7b320c fixed wrong isubs in IMP/Types
kleing
parents: 43150
diff changeset
    40
"tbval b1 s bv1 \<Longrightarrow> tbval b2 s bv2 \<Longrightarrow> tbval (And b1 b2) s (bv1 & bv2)" |
376c1e7b320c fixed wrong isubs in IMP/Types
kleing
parents: 43150
diff changeset
    41
"taval a1 s (Iv i1) \<Longrightarrow> taval a2 s (Iv i2) \<Longrightarrow> tbval (Less a1 a2) s (i1 < i2)" |
376c1e7b320c fixed wrong isubs in IMP/Types
kleing
parents: 43150
diff changeset
    42
"taval a1 s (Rv r1) \<Longrightarrow> taval a2 s (Rv r2) \<Longrightarrow> tbval (Less a1 a2) s (r1 < r2)"
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    43
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    44
subsection "Syntax of Commands"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    45
(* a copy of Com.thy - keep in sync! *)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    46
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    47
datatype
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    48
  com = SKIP 
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
    49
      | Assign vname aexp       ("_ ::= _" [1000, 61] 61)
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51461
diff changeset
    50
      | Seq    com  com         ("_;; _"  [60, 61] 60)
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    51
      | If     bexp com com     ("IF _ THEN _ ELSE _"  [0, 0, 61] 61)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    52
      | While  bexp com         ("WHILE _ DO _"  [0, 61] 61)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    53
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    54
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    55
subsection "Small-Step Semantics of Commands"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    56
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    57
inductive
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    58
  small_step :: "(com \<times> state) \<Rightarrow> (com \<times> state) \<Rightarrow> bool" (infix "\<rightarrow>" 55)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    59
where
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    60
Assign:  "taval a s v \<Longrightarrow> (x ::= a, s) \<rightarrow> (SKIP, s(x := v))" |
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    61
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51461
diff changeset
    62
Seq1:   "(SKIP;;c,s) \<rightarrow> (c,s)" |
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51461
diff changeset
    63
Seq2:   "(c1,s) \<rightarrow> (c1',s') \<Longrightarrow> (c1;;c2,s) \<rightarrow> (c1';;c2,s')" |
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    64
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    65
IfTrue:  "tbval b s True \<Longrightarrow> (IF b THEN c1 ELSE c2,s) \<rightarrow> (c1,s)" |
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    66
IfFalse: "tbval b s False \<Longrightarrow> (IF b THEN c1 ELSE c2,s) \<rightarrow> (c2,s)" |
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    67
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51461
diff changeset
    68
While:   "(WHILE b DO c,s) \<rightarrow> (IF b THEN c;; WHILE b DO c ELSE SKIP,s)"
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    69
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    70
lemmas small_step_induct = small_step.induct[split_format(complete)]
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    71
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    72
subsection "The Type System"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    73
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    74
datatype ty = Ity | Rty
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    75
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45200
diff changeset
    76
type_synonym tyenv = "vname \<Rightarrow> ty"
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    77
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    78
inductive atyping :: "tyenv \<Rightarrow> aexp \<Rightarrow> ty \<Rightarrow> bool"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    79
  ("(1_/ \<turnstile>/ (_ :/ _))" [50,0,50] 50)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    80
where
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    81
Ic_ty: "\<Gamma> \<turnstile> Ic i : Ity" |
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    82
Rc_ty: "\<Gamma> \<turnstile> Rc r : Rty" |
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    83
V_ty: "\<Gamma> \<turnstile> V x : \<Gamma> x" |
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    84
Plus_ty: "\<Gamma> \<turnstile> a1 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> a2 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> Plus a1 a2 : \<tau>"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    85
55572
fb3bb943a606 provide more automation for type definitions (makes book exercises easier)
kleing
parents: 54610
diff changeset
    86
declare atyping.intros [intro!]
fb3bb943a606 provide more automation for type definitions (makes book exercises easier)
kleing
parents: 54610
diff changeset
    87
inductive_cases [elim!]:
fb3bb943a606 provide more automation for type definitions (makes book exercises easier)
kleing
parents: 54610
diff changeset
    88
  "\<Gamma> \<turnstile> V x : \<tau>" "\<Gamma> \<turnstile> Ic i : \<tau>" "\<Gamma> \<turnstile> Rc r : \<tau>" "\<Gamma> \<turnstile> Plus a1 a2 : \<tau>"
fb3bb943a606 provide more automation for type definitions (makes book exercises easier)
kleing
parents: 54610
diff changeset
    89
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
    90
text\<open>Warning: the ``:'' notation leads to syntactic ambiguities,
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    91
i.e. multiple parse trees, because ``:'' also stands for set membership.
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    92
In most situations Isabelle's type system will reject all but one parse tree,
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
    93
but will still inform you of the potential ambiguity.\<close>
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    94
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    95
inductive btyping :: "tyenv \<Rightarrow> bexp \<Rightarrow> bool" (infix "\<turnstile>" 50)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    96
where
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
    97
B_ty: "\<Gamma> \<turnstile> Bc v" |
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    98
Not_ty: "\<Gamma> \<turnstile> b \<Longrightarrow> \<Gamma> \<turnstile> Not b" |
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
    99
And_ty: "\<Gamma> \<turnstile> b1 \<Longrightarrow> \<Gamma> \<turnstile> b2 \<Longrightarrow> \<Gamma> \<turnstile> And b1 b2" |
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   100
Less_ty: "\<Gamma> \<turnstile> a1 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> a2 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> Less a1 a2"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   101
55572
fb3bb943a606 provide more automation for type definitions (makes book exercises easier)
kleing
parents: 54610
diff changeset
   102
declare btyping.intros [intro!]
fb3bb943a606 provide more automation for type definitions (makes book exercises easier)
kleing
parents: 54610
diff changeset
   103
inductive_cases [elim!]: "\<Gamma> \<turnstile> Not b" "\<Gamma> \<turnstile> And b1 b2" "\<Gamma> \<turnstile> Less a1 a2"
fb3bb943a606 provide more automation for type definitions (makes book exercises easier)
kleing
parents: 54610
diff changeset
   104
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   105
inductive ctyping :: "tyenv \<Rightarrow> com \<Rightarrow> bool" (infix "\<turnstile>" 50) where
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   106
Skip_ty: "\<Gamma> \<turnstile> SKIP" |
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   107
Assign_ty: "\<Gamma> \<turnstile> a : \<Gamma>(x) \<Longrightarrow> \<Gamma> \<turnstile> x ::= a" |
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51461
diff changeset
   108
Seq_ty: "\<Gamma> \<turnstile> c1 \<Longrightarrow> \<Gamma> \<turnstile> c2 \<Longrightarrow> \<Gamma> \<turnstile> c1;;c2" |
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   109
If_ty: "\<Gamma> \<turnstile> b \<Longrightarrow> \<Gamma> \<turnstile> c1 \<Longrightarrow> \<Gamma> \<turnstile> c2 \<Longrightarrow> \<Gamma> \<turnstile> IF b THEN c1 ELSE c2" |
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   110
While_ty: "\<Gamma> \<turnstile> b \<Longrightarrow> \<Gamma> \<turnstile> c \<Longrightarrow> \<Gamma> \<turnstile> WHILE b DO c"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   111
55572
fb3bb943a606 provide more automation for type definitions (makes book exercises easier)
kleing
parents: 54610
diff changeset
   112
declare ctyping.intros [intro!]
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   113
inductive_cases [elim!]:
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51461
diff changeset
   114
  "\<Gamma> \<turnstile> x ::= a"  "\<Gamma> \<turnstile> c1;;c2"
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   115
  "\<Gamma> \<turnstile> IF b THEN c1 ELSE c2"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   116
  "\<Gamma> \<turnstile> WHILE b DO c"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   117
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   118
subsection "Well-typed Programs Do Not Get Stuck"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   119
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   120
fun type :: "val \<Rightarrow> ty" where
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   121
"type (Iv i) = Ity" |
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   122
"type (Rv r) = Rty"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   123
54610
nipkow
parents: 52405
diff changeset
   124
lemma type_eq_Ity[simp]: "type v = Ity \<longleftrightarrow> (\<exists>i. v = Iv i)"
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   125
by (cases v) simp_all
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   126
54610
nipkow
parents: 52405
diff changeset
   127
lemma type_eq_Rty[simp]: "type v = Rty \<longleftrightarrow> (\<exists>r. v = Rv r)"
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   128
by (cases v) simp_all
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   129
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   130
definition styping :: "tyenv \<Rightarrow> state \<Rightarrow> bool" (infix "\<turnstile>" 50)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   131
where "\<Gamma> \<turnstile> s  \<longleftrightarrow>  (\<forall>x. type (s x) = \<Gamma> x)"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   132
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   133
lemma apreservation:
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   134
  "\<Gamma> \<turnstile> a : \<tau> \<Longrightarrow> taval a s v \<Longrightarrow> \<Gamma> \<turnstile> s \<Longrightarrow> type v = \<tau>"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
   135
apply(induction arbitrary: v rule: atyping.induct)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44020
diff changeset
   136
apply (fastforce simp: styping_def)+
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   137
done
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   138
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   139
lemma aprogress: "\<Gamma> \<turnstile> a : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> s \<Longrightarrow> \<exists>v. taval a s v"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
   140
proof(induction rule: atyping.induct)
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   141
  case (Plus_ty \<Gamma> a1 t a2)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   142
  then obtain v1 v2 where v: "taval a1 s v1" "taval a2 s v2" by blast
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   143
  show ?case
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   144
  proof (cases v1)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   145
    case Iv
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
   146
    with Plus_ty v show ?thesis
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44020
diff changeset
   147
      by(fastforce intro: taval.intros(4) dest!: apreservation)
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   148
  next
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   149
    case Rv
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
   150
    with Plus_ty v show ?thesis
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44020
diff changeset
   151
      by(fastforce intro: taval.intros(5) dest!: apreservation)
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   152
  qed
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   153
qed (auto intro: taval.intros)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   154
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   155
lemma bprogress: "\<Gamma> \<turnstile> b \<Longrightarrow> \<Gamma> \<turnstile> s \<Longrightarrow> \<exists>v. tbval b s v"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
   156
proof(induction rule: btyping.induct)
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   157
  case (Less_ty \<Gamma> a1 t a2)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   158
  then obtain v1 v2 where v: "taval a1 s v1" "taval a2 s v2"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   159
    by (metis aprogress)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   160
  show ?case
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   161
  proof (cases v1)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   162
    case Iv
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   163
    with Less_ty v show ?thesis
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44020
diff changeset
   164
      by (fastforce intro!: tbval.intros(4) dest!:apreservation)
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   165
  next
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   166
    case Rv
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   167
    with Less_ty v show ?thesis
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44020
diff changeset
   168
      by (fastforce intro!: tbval.intros(5) dest!:apreservation)
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   169
  qed
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   170
qed (auto intro: tbval.intros)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   171
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   172
theorem progress:
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   173
  "\<Gamma> \<turnstile> c \<Longrightarrow> \<Gamma> \<turnstile> s \<Longrightarrow> c \<noteq> SKIP \<Longrightarrow> \<exists>cs'. (c,s) \<rightarrow> cs'"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
   174
proof(induction rule: ctyping.induct)
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   175
  case Skip_ty thus ?case by simp
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   176
next
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   177
  case Assign_ty 
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   178
  thus ?case by (metis Assign aprogress)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   179
next
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   180
  case Seq_ty thus ?case by simp (metis Seq1 Seq2)
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   181
next
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   182
  case (If_ty \<Gamma> b c1 c2)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   183
  then obtain bv where "tbval b s bv" by (metis bprogress)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   184
  show ?case
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   185
  proof(cases bv)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   186
    assume "bv"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   187
    with \<open>tbval b s bv\<close> show ?case by simp (metis IfTrue)
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   188
  next
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   189
    assume "\<not>bv"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   190
    with \<open>tbval b s bv\<close> show ?case by simp (metis IfFalse)
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   191
  qed
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   192
next
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   193
  case While_ty show ?case by (metis While)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   194
qed
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   195
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   196
theorem styping_preservation:
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   197
  "(c,s) \<rightarrow> (c',s') \<Longrightarrow> \<Gamma> \<turnstile> c \<Longrightarrow> \<Gamma> \<turnstile> s \<Longrightarrow> \<Gamma> \<turnstile> s'"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
   198
proof(induction rule: small_step_induct)
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   199
  case Assign thus ?case
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   200
    by (auto simp: styping_def) (metis Assign(1,3) apreservation)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   201
qed auto
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   202
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   203
theorem ctyping_preservation:
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   204
  "(c,s) \<rightarrow> (c',s') \<Longrightarrow> \<Gamma> \<turnstile> c \<Longrightarrow> \<Gamma> \<turnstile> c'"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   205
by (induct rule: small_step_induct) (auto simp: ctyping.intros)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   206
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   207
abbreviation small_steps :: "com * state \<Rightarrow> com * state \<Rightarrow> bool" (infix "\<rightarrow>*" 55)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   208
where "x \<rightarrow>* y == star small_step x y"
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   209
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   210
theorem type_sound:
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   211
  "(c,s) \<rightarrow>* (c',s') \<Longrightarrow> \<Gamma> \<turnstile> c \<Longrightarrow> \<Gamma> \<turnstile> s \<Longrightarrow> c' \<noteq> SKIP
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   212
   \<Longrightarrow> \<exists>cs''. (c',s') \<rightarrow> cs''"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
   213
apply(induction rule:star_induct)
43150
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   214
apply (metis progress)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   215
by (metis styping_preservation ctyping_preservation)
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   216
69bc4dafcc53 Added typed IMP
nipkow
parents:
diff changeset
   217
end