src/HOL/IMP/Types.thy
changeset 51454 9ac7868ae64f
parent 51394 27bb849330ea
child 51461 e1e8191c6725
equal deleted inserted replaced
51453:d2bc229e1f37 51454:9ac7868ae64f
     9 type_synonym vname = string
     9 type_synonym vname = string
    10 type_synonym state = "vname \<Rightarrow> val"
    10 type_synonym state = "vname \<Rightarrow> val"
    11 
    11 
    12 datatype aexp =  Ic int | Rc real | V vname | Plus aexp aexp
    12 datatype aexp =  Ic int | Rc real | V vname | Plus aexp aexp
    13 
    13 
    14 text_raw{*\snip{tavalDef}{0}{2}{% *}
       
    15 inductive taval :: "aexp \<Rightarrow> state \<Rightarrow> val \<Rightarrow> bool" where
    14 inductive taval :: "aexp \<Rightarrow> state \<Rightarrow> val \<Rightarrow> bool" where
    16 "taval (Ic i) s (Iv i)" |
    15 "taval (Ic i) s (Iv i)" |
    17 "taval (Rc r) s (Rv r)" |
    16 "taval (Rc r) s (Rv r)" |
    18 "taval (V x) s (s x)" |
    17 "taval (V x) s (s x)" |
    19 "taval a1 s (Iv i1) \<Longrightarrow> taval a2 s (Iv i2)
    18 "taval a1 s (Iv i1) \<Longrightarrow> taval a2 s (Iv i2)
    20  \<Longrightarrow> taval (Plus a1 a2) s (Iv(i1+i2))" |
    19  \<Longrightarrow> taval (Plus a1 a2) s (Iv(i1+i2))" |
    21 "taval a1 s (Rv r1) \<Longrightarrow> taval a2 s (Rv r2)
    20 "taval a1 s (Rv r1) \<Longrightarrow> taval a2 s (Rv r2)
    22  \<Longrightarrow> taval (Plus a1 a2) s (Rv(r1+r2))"
    21  \<Longrightarrow> taval (Plus a1 a2) s (Rv(r1+r2))"
    23 text_raw{*}%endsnip*}
       
    24 
    22 
    25 inductive_cases [elim!]:
    23 inductive_cases [elim!]:
    26   "taval (Ic i) s v"  "taval (Rc i) s v"
    24   "taval (Ic i) s v"  "taval (Rc i) s v"
    27   "taval (V x) s v"
    25   "taval (V x) s v"
    28   "taval (Plus a1 a2) s v"
    26   "taval (Plus a1 a2) s v"
    29 
    27 
    30 subsection "Boolean Expressions"
    28 subsection "Boolean Expressions"
    31 
    29 
    32 datatype bexp = Bc bool | Not bexp | And bexp bexp | Less aexp aexp
    30 datatype bexp = Bc bool | Not bexp | And bexp bexp | Less aexp aexp
    33 
    31 
    34 text_raw{*\snip{tbvalDef}{0}{2}{% *}
       
    35 inductive tbval :: "bexp \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> bool" where
    32 inductive tbval :: "bexp \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> bool" where
    36 "tbval (Bc v) s v" |
    33 "tbval (Bc v) s v" |
    37 "tbval b s bv \<Longrightarrow> tbval (Not b) s (\<not> bv)" |
    34 "tbval b s bv \<Longrightarrow> tbval (Not b) s (\<not> bv)" |
    38 "tbval b1 s bv1 \<Longrightarrow> tbval b2 s bv2 \<Longrightarrow> tbval (And b1 b2) s (bv1 & bv2)" |
    35 "tbval b1 s bv1 \<Longrightarrow> tbval b2 s bv2 \<Longrightarrow> tbval (And b1 b2) s (bv1 & bv2)" |
    39 "taval a1 s (Iv i1) \<Longrightarrow> taval a2 s (Iv i2) \<Longrightarrow> tbval (Less a1 a2) s (i1 < i2)" |
    36 "taval a1 s (Iv i1) \<Longrightarrow> taval a2 s (Iv i2) \<Longrightarrow> tbval (Less a1 a2) s (i1 < i2)" |
    40 "taval a1 s (Rv r1) \<Longrightarrow> taval a2 s (Rv r2) \<Longrightarrow> tbval (Less a1 a2) s (r1 < r2)"
    37 "taval a1 s (Rv r1) \<Longrightarrow> taval a2 s (Rv r2) \<Longrightarrow> tbval (Less a1 a2) s (r1 < r2)"
    41 text_raw{*}%endsnip*}
       
    42 
    38 
    43 subsection "Syntax of Commands"
    39 subsection "Syntax of Commands"
    44 (* a copy of Com.thy - keep in sync! *)
    40 (* a copy of Com.thy - keep in sync! *)
    45 
    41 
    46 datatype
    42 datatype
    51       | While  bexp com         ("WHILE _ DO _"  [0, 61] 61)
    47       | While  bexp com         ("WHILE _ DO _"  [0, 61] 61)
    52 
    48 
    53 
    49 
    54 subsection "Small-Step Semantics of Commands"
    50 subsection "Small-Step Semantics of Commands"
    55 
    51 
    56 text_raw{*\snip{tsmallstepDef}{0}{2}{% *}
       
    57 inductive
    52 inductive
    58   small_step :: "(com \<times> state) \<Rightarrow> (com \<times> state) \<Rightarrow> bool" (infix "\<rightarrow>" 55)
    53   small_step :: "(com \<times> state) \<Rightarrow> (com \<times> state) \<Rightarrow> bool" (infix "\<rightarrow>" 55)
    59 where
    54 where
    60 Assign:  "taval a s v \<Longrightarrow> (x ::= a, s) \<rightarrow> (SKIP, s(x := v))" |
    55 Assign:  "taval a s v \<Longrightarrow> (x ::= a, s) \<rightarrow> (SKIP, s(x := v))" |
    61 
    56 
    64 
    59 
    65 IfTrue:  "tbval b s True \<Longrightarrow> (IF b THEN c1 ELSE c2,s) \<rightarrow> (c1,s)" |
    60 IfTrue:  "tbval b s True \<Longrightarrow> (IF b THEN c1 ELSE c2,s) \<rightarrow> (c1,s)" |
    66 IfFalse: "tbval b s False \<Longrightarrow> (IF b THEN c1 ELSE c2,s) \<rightarrow> (c2,s)" |
    61 IfFalse: "tbval b s False \<Longrightarrow> (IF b THEN c1 ELSE c2,s) \<rightarrow> (c2,s)" |
    67 
    62 
    68 While:   "(WHILE b DO c,s) \<rightarrow> (IF b THEN c; WHILE b DO c ELSE SKIP,s)"
    63 While:   "(WHILE b DO c,s) \<rightarrow> (IF b THEN c; WHILE b DO c ELSE SKIP,s)"
    69 text_raw{*}%endsnip*}
       
    70 
    64 
    71 lemmas small_step_induct = small_step.induct[split_format(complete)]
    65 lemmas small_step_induct = small_step.induct[split_format(complete)]
    72 
    66 
    73 subsection "The Type System"
    67 subsection "The Type System"
    74 
    68 
    75 datatype ty = Ity | Rty
    69 datatype ty = Ity | Rty
    76 
    70 
    77 type_synonym tyenv = "vname \<Rightarrow> ty"
    71 type_synonym tyenv = "vname \<Rightarrow> ty"
    78 
    72 
    79 text_raw{*\snip{atypingDef}{0}{2}{% *}
       
    80 inductive atyping :: "tyenv \<Rightarrow> aexp \<Rightarrow> ty \<Rightarrow> bool"
    73 inductive atyping :: "tyenv \<Rightarrow> aexp \<Rightarrow> ty \<Rightarrow> bool"
    81   ("(1_/ \<turnstile>/ (_ :/ _))" [50,0,50] 50)
    74   ("(1_/ \<turnstile>/ (_ :/ _))" [50,0,50] 50)
    82 where
    75 where
    83 Ic_ty: "\<Gamma> \<turnstile> Ic i : Ity" |
    76 Ic_ty: "\<Gamma> \<turnstile> Ic i : Ity" |
    84 Rc_ty: "\<Gamma> \<turnstile> Rc r : Rty" |
    77 Rc_ty: "\<Gamma> \<turnstile> Rc r : Rty" |
    85 V_ty: "\<Gamma> \<turnstile> V x : \<Gamma> x" |
    78 V_ty: "\<Gamma> \<turnstile> V x : \<Gamma> x" |
    86 Plus_ty: "\<Gamma> \<turnstile> a1 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> a2 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> Plus a1 a2 : \<tau>"
    79 Plus_ty: "\<Gamma> \<turnstile> a1 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> a2 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> Plus a1 a2 : \<tau>"
    87 text_raw{*}%endsnip*}
       
    88 
    80 
    89 text{* Warning: the ``:'' notation leads to syntactic ambiguities,
    81 text{* Warning: the ``:'' notation leads to syntactic ambiguities,
    90 i.e. multiple parse trees, because ``:'' also stands for set membership.
    82 i.e. multiple parse trees, because ``:'' also stands for set membership.
    91 In most situations Isabelle's type system will reject all but one parse tree,
    83 In most situations Isabelle's type system will reject all but one parse tree,
    92 but will still inform you of the potential ambiguity. *}
    84 but will still inform you of the potential ambiguity. *}
    93 
    85 
    94 text_raw{*\snip{btypingDef}{0}{2}{% *}
       
    95 inductive btyping :: "tyenv \<Rightarrow> bexp \<Rightarrow> bool" (infix "\<turnstile>" 50)
    86 inductive btyping :: "tyenv \<Rightarrow> bexp \<Rightarrow> bool" (infix "\<turnstile>" 50)
    96 where
    87 where
    97 B_ty: "\<Gamma> \<turnstile> Bc v" |
    88 B_ty: "\<Gamma> \<turnstile> Bc v" |
    98 Not_ty: "\<Gamma> \<turnstile> b \<Longrightarrow> \<Gamma> \<turnstile> Not b" |
    89 Not_ty: "\<Gamma> \<turnstile> b \<Longrightarrow> \<Gamma> \<turnstile> Not b" |
    99 And_ty: "\<Gamma> \<turnstile> b1 \<Longrightarrow> \<Gamma> \<turnstile> b2 \<Longrightarrow> \<Gamma> \<turnstile> And b1 b2" |
    90 And_ty: "\<Gamma> \<turnstile> b1 \<Longrightarrow> \<Gamma> \<turnstile> b2 \<Longrightarrow> \<Gamma> \<turnstile> And b1 b2" |
   100 Less_ty: "\<Gamma> \<turnstile> a1 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> a2 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> Less a1 a2"
    91 Less_ty: "\<Gamma> \<turnstile> a1 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> a2 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> Less a1 a2"
   101 text_raw{*}%endsnip*}
    92 
   102 
       
   103 text_raw{*\snip{ctypingDef}{0}{2}{% *}
       
   104 inductive ctyping :: "tyenv \<Rightarrow> com \<Rightarrow> bool" (infix "\<turnstile>" 50) where
    93 inductive ctyping :: "tyenv \<Rightarrow> com \<Rightarrow> bool" (infix "\<turnstile>" 50) where
   105 Skip_ty: "\<Gamma> \<turnstile> SKIP" |
    94 Skip_ty: "\<Gamma> \<turnstile> SKIP" |
   106 Assign_ty: "\<Gamma> \<turnstile> a : \<Gamma>(x) \<Longrightarrow> \<Gamma> \<turnstile> x ::= a" |
    95 Assign_ty: "\<Gamma> \<turnstile> a : \<Gamma>(x) \<Longrightarrow> \<Gamma> \<turnstile> x ::= a" |
   107 Seq_ty: "\<Gamma> \<turnstile> c1 \<Longrightarrow> \<Gamma> \<turnstile> c2 \<Longrightarrow> \<Gamma> \<turnstile> c1;c2" |
    96 Seq_ty: "\<Gamma> \<turnstile> c1 \<Longrightarrow> \<Gamma> \<turnstile> c2 \<Longrightarrow> \<Gamma> \<turnstile> c1;c2" |
   108 If_ty: "\<Gamma> \<turnstile> b \<Longrightarrow> \<Gamma> \<turnstile> c1 \<Longrightarrow> \<Gamma> \<turnstile> c2 \<Longrightarrow> \<Gamma> \<turnstile> IF b THEN c1 ELSE c2" |
    97 If_ty: "\<Gamma> \<turnstile> b \<Longrightarrow> \<Gamma> \<turnstile> c1 \<Longrightarrow> \<Gamma> \<turnstile> c2 \<Longrightarrow> \<Gamma> \<turnstile> IF b THEN c1 ELSE c2" |
   109 While_ty: "\<Gamma> \<turnstile> b \<Longrightarrow> \<Gamma> \<turnstile> c \<Longrightarrow> \<Gamma> \<turnstile> WHILE b DO c"
    98 While_ty: "\<Gamma> \<turnstile> b \<Longrightarrow> \<Gamma> \<turnstile> c \<Longrightarrow> \<Gamma> \<turnstile> WHILE b DO c"
   110 text_raw{*}%endsnip*}
       
   111 
    99 
   112 inductive_cases [elim!]:
   100 inductive_cases [elim!]:
   113   "\<Gamma> \<turnstile> x ::= a"  "\<Gamma> \<turnstile> c1;c2"
   101   "\<Gamma> \<turnstile> x ::= a"  "\<Gamma> \<turnstile> c1;c2"
   114   "\<Gamma> \<turnstile> IF b THEN c1 ELSE c2"
   102   "\<Gamma> \<turnstile> IF b THEN c1 ELSE c2"
   115   "\<Gamma> \<turnstile> WHILE b DO c"
   103   "\<Gamma> \<turnstile> WHILE b DO c"