src/HOL/Datatype_Examples/Milner_Tofte.thy
author traytel
Tue, 12 Jan 2016 09:28:08 +0100
changeset 62136 c92d82c3f41b
parent 61745 e23e0ff98657
permissions -rw-r--r--
removed outdated example
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61745
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
     1
(*  Title:      HOL/Datatype_Examples/Milner_Tofte.thy
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
     2
    Author:     Dmitriy Traytel, ETH Zürich
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
     3
    Copyright   2015
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
     4
62136
c92d82c3f41b removed outdated example
traytel
parents: 61745
diff changeset
     5
Modernized version of an old development by Jacob Frost
61745
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
     6
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
     7
Based upon the article
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
     8
    Robin Milner and Mads Tofte,
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
     9
    Co-induction in Relational Semantics,
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    10
    Theoretical Computer Science 87 (1991), pages 209-220.
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    11
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    12
Written up as
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    13
    Jacob Frost, A Case Study of Co_induction in Isabelle/HOL
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    14
    Report 308, Computer Lab, University of Cambridge (1993).
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    15
*)
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    16
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    17
section \<open>Milner-Tofte: Co-induction in Relational Semantics\<close>
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    18
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    19
theory Milner_Tofte
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    20
imports Main
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    21
begin
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    22
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    23
typedecl Const
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    24
typedecl ExVar
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    25
typedecl TyConst
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    26
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    27
datatype Ex =
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    28
  e_const (e_const_fst: Const)
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    29
| e_var ExVar
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    30
| e_fn ExVar Ex ("fn _ => _" [0,51] 1000)
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    31
| e_fix ExVar ExVar Ex ("fix _ ( _ ) = _" [0,51,51] 1000)
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    32
| e_app Ex Ex ("_ @@ _" [51,51] 1000)
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    33
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    34
datatype Ty =
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    35
  t_const TyConst
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    36
| t_fun Ty Ty ("_ -> _" [51,51] 1000)
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    37
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    38
datatype 'a genClos =
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    39
  clos_mk ExVar Ex "ExVar \<rightharpoonup> 'a" ("\<langle>_ , _ , _\<rangle>" [0,0,0] 1000)
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    40
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    41
codatatype Val =
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    42
  v_const Const
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    43
| v_clos "Val genClos"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    44
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    45
type_synonym Clos = "Val genClos"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    46
type_synonym ValEnv = "ExVar \<rightharpoonup> Val"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    47
type_synonym TyEnv = "ExVar \<rightharpoonup> Ty"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    48
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    49
axiomatization 
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    50
  c_app :: "[Const, Const] => Const" and
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    51
  isof :: "[Const, Ty] => bool" ("_ isof _" [36,36] 50) where
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    52
  isof_app: "\<lbrakk>c1 isof t1 -> t2; c2 isof t1\<rbrakk> \<Longrightarrow> c_app c1 c2 isof t2"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    53
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    54
text \<open>The dynamic semantics is defined inductively by a set of inference
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    55
rules.  These inference rules allows one to draw conclusions of the form ve
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    56
|- e ---> v, read the expression e evaluates to the value v in the value
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    57
environment ve.  Therefore the relation _ |- _ ---> _ is defined in Isabelle
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    58
as the least fixpoint of the functor eval_fun below.  From this definition
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    59
introduction rules and a strong elimination (induction) rule can be derived.\<close>
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    60
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    61
inductive eval :: "[ValEnv, Ex, Val] => bool" ("_ \<turnstile> _ ---> _" [36,0,36] 50) where
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    62
  eval_const: "ve \<turnstile> e_const c ---> v_const c"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    63
| eval_var2:  "ev \<in> dom ve  \<Longrightarrow> ve \<turnstile> e_var ev ---> the (ve ev)"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    64
| eval_fn:    "ve \<turnstile> fn ev => e ---> v_clos \<langle>ev, e, ve\<rangle>"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    65
| eval_fix:   "cl = \<langle>ev1, e, ve(ev2 \<mapsto> v_clos cl)\<rangle> \<Longrightarrow> ve \<turnstile> fix ev2(ev1) = e ---> v_clos(cl)"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    66
| eval_app1:  "\<lbrakk>ve \<turnstile> e1 ---> v_const c1; ve \<turnstile> e2 ---> v_const c2\<rbrakk> \<Longrightarrow>
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    67
                ve \<turnstile> e1 @@ e2 ---> v_const (c_app c1 c2)"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    68
| eval_app2:  "\<lbrakk>ve \<turnstile> e1 ---> v_clos \<langle>xm, em, vem\<rangle>; ve \<turnstile> e2 ---> v2; vem(xm \<mapsto> v2) \<turnstile> em ---> v\<rbrakk> \<Longrightarrow>
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    69
                ve \<turnstile> e1 @@ e2 ---> v"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    70
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    71
declare eval.intros[intro]
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    72
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    73
text \<open>The static semantics is defined in the same way as the dynamic
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    74
semantics.  The relation te |- e ===> t express the expression e has the
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    75
type t in the type environment te.\<close>
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    76
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    77
inductive elab :: "[TyEnv, Ex, Ty] => bool" ("_ \<turnstile> _ ===> _" [36,0,36] 50) where
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    78
  elab_const: "c isof ty \<Longrightarrow> te \<turnstile> e_const c ===> ty"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    79
| elab_var:   "x \<in> dom te \<Longrightarrow> te \<turnstile> e_var x ===> the (te x)"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    80
| elab_fn:    "te(x \<mapsto> ty1) \<turnstile> e ===> ty2 \<Longrightarrow> te \<turnstile> fn x => e ===> ty1 -> ty2"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    81
| elab_fix:   "te(f \<mapsto> ty1 -> ty2, x \<mapsto> ty1) \<turnstile> e ===> ty2 \<Longrightarrow> te \<turnstile> fix f x = e ===> ty1 -> ty2"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    82
| elab_app:   "\<lbrakk>te \<turnstile> e1 ===> ty1 -> ty2; te \<turnstile> e2 ===> ty1\<rbrakk> \<Longrightarrow> te \<turnstile> e1 @@ e2 ===> ty2"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    83
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    84
declare elab.intros[intro]
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    85
inductive_cases elabE[elim!]:
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    86
  "te \<turnstile> e_const(c) ===> ty"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    87
  "te \<turnstile> e_var(x) ===> ty"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    88
  "te \<turnstile> fn x => e ===> ty"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    89
  "te \<turnstile> fix f(x) = e ===> ty"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    90
  "te \<turnstile> e1 @@ e2 ===> ty"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    91
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    92
(* The original correspondence relation *)
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    93
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    94
abbreviation isof_env :: "[ValEnv,TyEnv] => bool" ("_ isofenv _") where
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    95
  "ve isofenv te \<equiv> (dom(ve) = dom(te) \<and>
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    96
     (\<forall>x. x \<in> dom ve \<longrightarrow> (\<exists>c. the (ve x) = v_const(c) \<and> c isof the (te x))))"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    97
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    98
coinductive hasty :: "[Val, Ty] => bool" ("_ hasty _" [36,36] 50) where
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
    99
  hasty_const: "c isof t \<Longrightarrow> v_const c hasty t"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   100
| hasty_clos:  "\<lbrakk>te \<turnstile> fn ev => e ===> t; dom(ve) = dom(te) \<and>
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   101
   (\<forall>x. x \<in> dom ve --> the (ve x) hasty the (te x)); cl = \<langle>ev,e,ve\<rangle>\<rbrakk> \<Longrightarrow> v_clos cl hasty t"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   102
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   103
declare hasty.intros[intro]
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   104
inductive_cases hastyE[elim!]:
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   105
  "v_const c hasty t"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   106
  "v_clos \<langle>xm , em , evm\<rangle> hasty u -> t"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   107
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   108
definition hasty_env :: "[ValEnv, TyEnv] => bool" ("_ hastyenv _ " [36,36] 35) where
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   109
  [simp]: "ve hastyenv te \<equiv> (dom(ve) = dom(te) \<and>
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   110
     (\<forall>x. x \<in> dom ve --> the (ve x) hasty the (te x)))"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   111
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   112
theorem consistency: "\<lbrakk>ve \<turnstile> e ---> v; ve hastyenv te; te \<turnstile> e ===> t\<rbrakk> \<Longrightarrow> v hasty t"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   113
proof (induct ve e v arbitrary: t te rule: eval.induct)
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   114
  case (eval_fix cl x e ve f)
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   115
  then show ?case
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   116
    by coinduction
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   117
      (auto 0 11 intro: exI[of _ "te(f \<mapsto> _)"] exI[of _ x] exI[of _ e] exI[of _ "ve(f \<mapsto> v_clos cl)"])
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   118
next
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   119
  case (eval_app2 ve e1 xm em evm e2 v2 v)
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   120
  then obtain u where "te \<turnstile> e1 ===> u -> t" "te \<turnstile> e2 ===> u" by auto
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   121
  with eval_app2(2)[of te "u -> t"] eval_app2(4)[of te u] eval_app2(1,3,5,7) show ?case
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   122
    by (auto 0 4 elim!: eval_app2(6)[rotated])
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   123
qed (auto 8 0 intro!: isof_app)
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   124
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   125
lemma basic_consistency_aux:
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   126
  "ve isofenv te \<Longrightarrow> ve hastyenv te"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   127
  using hasty_const hasty_env_def by force
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   128
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   129
theorem basic_consistency:
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   130
  "\<lbrakk>ve isofenv te; ve \<turnstile> e ---> v_const c; te \<turnstile> e ===> t\<rbrakk> \<Longrightarrow> c isof t"
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   131
  by (auto dest: consistency intro!: basic_consistency_aux)
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   132
e23e0ff98657 Ported old example to use (co)datatypes
traytel
parents:
diff changeset
   133
end