src/HOL/Nominal/Examples/Lambda_mu.thy
author urbanc
Mon, 07 Nov 2005 15:19:03 +0100
changeset 18106 836135c8acb2
child 18269 3f36e2165e51
permissions -rw-r--r--
Initial commit.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     1
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     2
theory lambda_mu 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     3
imports "../nominal" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     4
begin
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     5
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     6
section {* Mu-Calculus from Gavin's cilmu-Paper*}
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     7
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     8
atom_decl var mvar
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     9
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    10
nominal_datatype trm = Var   "var" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    11
                     | Lam  "\<guillemotleft>var\<guillemotright>trm"   ("Lam [_]._" [100,100] 100)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    12
                     | App  "trm" "trm" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    13
                     | Pss  "mvar" "trm"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    14
                     | Act  "\<guillemotleft>mvar\<guillemotright>trm"  ("Act [_]._" [100,100] 100)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    15
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    16
section {* strong induction principle *}
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    17
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    18
lemma trm_induct_aux:
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    19
  fixes P :: "trm \<Rightarrow> 'a \<Rightarrow> bool"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    20
  and   f1 :: "'a \<Rightarrow> var set"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    21
  and   f2 :: "'a \<Rightarrow> mvar set"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    22
  assumes fs1: "\<And>x. finite (f1 x)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    23
      and fs2: "\<And>x. finite (f2 x)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    24
      and h1: "\<And>k x. P (Var x) k"  
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    25
      and h2: "\<And>k x t. x\<notin>f1 k \<Longrightarrow> (\<forall>l. P t l) \<Longrightarrow> P (Lam [x].t) k" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    26
      and h3: "\<And>k t1 t2. (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l) \<Longrightarrow> P (App t1 t2) k" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    27
      and h4: "\<And>k a t1. (\<forall>l. P t1 l) \<Longrightarrow> P (Pss a t1) k" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    28
      and h5: "\<And>k a t1. a\<notin>f2 k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> P (Act [a].t1) k"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    29
  shows "\<forall>(pi1::var prm) (pi2::mvar prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    30
proof (induct rule: trm.induct_weak)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    31
  case (goal1 a)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    32
  show ?case using h1 by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    33
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    34
  case (goal2 x t)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    35
  assume i1: "\<forall>(pi1::var prm)(pi2::mvar prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    36
  show ?case
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    37
  proof (intro strip, simp add: abs_perm)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    38
    fix pi1::"var prm" and pi2::"mvar prm" and k::"'a"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    39
    have f: "\<exists>c::var. c\<sharp>(f1 k,pi1\<bullet>(pi2\<bullet>x),pi1\<bullet>(pi2\<bullet>t))"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    40
      by (rule at_exists_fresh[OF at_var_inst], simp add: supp_prod fs_var1 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    41
          at_fin_set_supp[OF at_var_inst, OF fs1] fs1)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    42
    then obtain c::"var" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    43
      where f1: "c\<noteq>(pi1\<bullet>(pi2\<bullet>x))" and f2: "c\<sharp>(f1 k)" and f3: "c\<sharp>(pi1\<bullet>(pi2\<bullet>t))" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    44
      by (force simp add: fresh_prod at_fresh[OF at_var_inst])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    45
    have g: "Lam [c].([(c,pi1\<bullet>(pi2\<bullet>x))]\<bullet>(pi1\<bullet>(pi2\<bullet>t))) = Lam [(pi1\<bullet>(pi2\<bullet>x))].(pi1\<bullet>(pi2\<bullet>t))" using f1 f3
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    46
      by (simp add: trm.inject alpha)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    47
    from i1 have "\<forall>k. P (([(c,pi1\<bullet>(pi2\<bullet>x))]@pi1)\<bullet>(pi2\<bullet>t)) k" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    48
    hence i1b: "\<forall>k. P ([(c,pi1\<bullet>(pi2\<bullet>x))]\<bullet>(pi1\<bullet>(pi2\<bullet>t))) k" by (simp add: pt_var2[symmetric])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    49
    with h3 f2 have "P (Lam [c].([(c,pi1\<bullet>(pi2\<bullet>x))]\<bullet>(pi1\<bullet>(pi2\<bullet>t)))) k" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    50
      by (auto simp add: fresh_def at_fin_set_supp[OF at_var_inst, OF fs1])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    51
    with g show "P (Lam [(pi1\<bullet>(pi2\<bullet>x))].(pi1\<bullet>(pi2\<bullet>t))) k" by simp 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    52
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    53
next 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    54
  case (goal3 t1 t2)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    55
  assume i1: "\<forall>(pi1::var prm)(pi2::mvar prm) k. P (pi1\<bullet>(pi2\<bullet>t1)) k" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    56
  assume i2: "\<forall>(pi1::var prm)(pi2::mvar prm) k. P (pi1\<bullet>(pi2\<bullet>t2)) k"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    57
  show ?case 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    58
  proof (intro strip)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    59
    fix pi1::"var prm" and pi2::"mvar prm" and k::"'a"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    60
    from h3 i1 i2 have "P (App (pi1\<bullet>(pi2\<bullet>t1)) (pi1\<bullet>(pi2\<bullet>t2))) k" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    61
    thus "P (pi1\<bullet>(pi2\<bullet>(App t1 t2))) k" by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    62
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    63
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    64
  case (goal4 b t)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    65
  assume i1: "\<forall>(pi1::var prm)(pi2::mvar prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    66
  show ?case 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    67
  proof (intro strip)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    68
    fix pi1::"var prm" and pi2::"mvar prm" and k::"'a"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    69
    from h4 i1 have "P (Pss (pi1\<bullet>(pi2\<bullet>b)) (pi1\<bullet>(pi2\<bullet>t))) k" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    70
    thus "P (pi1\<bullet>(pi2\<bullet>(Pss b t))) k" by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    71
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    72
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    73
  case (goal5 b t)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    74
  assume i1: "\<forall>(pi1::var prm)(pi2::mvar prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    75
  show ?case
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    76
  proof (intro strip, simp add: abs_perm)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    77
    fix pi1::"var prm" and pi2::"mvar prm" and k::"'a"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    78
    have f: "\<exists>c::mvar. c\<sharp>(f2 k,pi1\<bullet>(pi2\<bullet>b),pi1\<bullet>(pi2\<bullet>t))"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    79
      by (rule at_exists_fresh[OF at_mvar_inst], simp add: supp_prod fs_mvar1 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    80
          at_fin_set_supp[OF at_mvar_inst, OF fs2] fs2)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    81
    then obtain c::"mvar" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    82
      where f1: "c\<noteq>(pi1\<bullet>(pi2\<bullet>b))" and f2: "c\<sharp>(f2 k)" and f3: "c\<sharp>(pi1\<bullet>(pi2\<bullet>t))" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    83
      by (force simp add: fresh_prod at_fresh[OF at_mvar_inst])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    84
    have g: "Act [c].(pi1\<bullet>([(c,pi1\<bullet>(pi2\<bullet>b))]\<bullet>(pi2\<bullet>t))) = Act [(pi1\<bullet>(pi2\<bullet>b))].(pi1\<bullet>(pi2\<bullet>t))" using f1 f3
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    85
      by (simp add: trm.inject alpha, simp add: dj_cp[OF cp_mvar_var_inst, OF dj_var_mvar])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    86
    from i1 have "\<forall>k. P (pi1\<bullet>(([(c,pi1\<bullet>(pi2\<bullet>b))]@pi2)\<bullet>t)) k" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    87
    hence i1b: "\<forall>k. P (pi1\<bullet>([(c,pi1\<bullet>(pi2\<bullet>b))]\<bullet>(pi2\<bullet>t))) k" by (simp add: pt_mvar2[symmetric])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    88
    with h5 f2 have "P (Act [c].(pi1\<bullet>([(c,pi1\<bullet>(pi2\<bullet>b))]\<bullet>(pi2\<bullet>t)))) k" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    89
      by (auto simp add: fresh_def at_fin_set_supp[OF at_mvar_inst, OF fs2])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    90
    with g show "P (Act [(pi1\<bullet>(pi2\<bullet>b))].(pi1\<bullet>(pi2\<bullet>t))) k" by simp 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    91
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    92
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    93
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    94
lemma trm_induct'[case_names Var Lam App Pss Act]:
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    95
  fixes P :: "trm \<Rightarrow> 'a \<Rightarrow> bool"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    96
  and   f1 :: "'a \<Rightarrow> var set"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    97
  and   f2 :: "'a \<Rightarrow> mvar set"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    98
  assumes fs1: "\<And>x. finite (f1 x)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    99
      and fs2: "\<And>x. finite (f2 x)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   100
      and h1: "\<And>k x. P (Var x) k"  
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   101
      and h2: "\<And>k x t. x\<notin>f1 k \<Longrightarrow> (\<forall>l. P t l) \<Longrightarrow> P (Lam [x].t) k" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   102
      and h3: "\<And>k t1 t2. (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l) \<Longrightarrow> P (App t1 t2) k" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   103
      and h4: "\<And>k a t1. (\<forall>l. P t1 l) \<Longrightarrow> P (Pss a t1) k" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   104
      and h5: "\<And>k a t1. a\<notin>f2 k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> P (Act [a].t1) k"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   105
  shows "P t k"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   106
proof -
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   107
  have "\<forall>(pi1::var prm)(pi2::mvar prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   108
  using fs1 fs2 h1 h2 h3 h4 h5 by (rule trm_induct_aux, auto)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   109
  hence "P (([]::var prm)\<bullet>(([]::mvar prm)\<bullet>t)) k" by blast
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   110
  thus "P t k" by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   111
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   112
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   113
lemma trm_induct[case_names Var Lam App Pss Act]: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   114
  fixes P :: "trm \<Rightarrow> ('a::{fs_var,fs_mvar}) \<Rightarrow> bool"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   115
  assumes h1: "\<And>k x. P (Var x) k"  
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   116
      and h2: "\<And>k x t. x\<sharp>k \<Longrightarrow> (\<forall>l. P t l) \<Longrightarrow> P (Lam [x].t) k" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   117
      and h3: "\<And>k t1 t2. (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l) \<Longrightarrow> P (App t1 t2) k" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   118
      and h4: "\<And>k a t1. (\<forall>l. P t1 l) \<Longrightarrow> P (Pss a t1) k" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   119
      and h5: "\<And>k a t1. a\<sharp>k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> P (Act [a].t1) k"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   120
  shows  "P t k"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   121
by (rule trm_induct'[of "\<lambda>x. ((supp x)::var set)" "\<lambda>x. ((supp x)::mvar set)" "P"], 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   122
    simp_all add: fs_var1 fs_mvar1 fresh_def[symmetric], auto intro: h1 h2 h3 h4 h5)