src/HOL/Nominal/Examples/Lambda_mu.thy
changeset 18659 2ff0ae57431d
parent 18269 3f36e2165e51
child 19501 9afa7183dfc2
equal deleted inserted replaced
18658:317a6f0ef8b9 18659:2ff0ae57431d
    12                      | Lam  "\<guillemotleft>var\<guillemotright>trm"   ("Lam [_]._" [100,100] 100)
    12                      | Lam  "\<guillemotleft>var\<guillemotright>trm"   ("Lam [_]._" [100,100] 100)
    13                      | App  "trm" "trm" 
    13                      | App  "trm" "trm" 
    14                      | Pss  "mvar" "trm"
    14                      | Pss  "mvar" "trm"
    15                      | Act  "\<guillemotleft>mvar\<guillemotright>trm"  ("Act [_]._" [100,100] 100)
    15                      | Act  "\<guillemotleft>mvar\<guillemotright>trm"  ("Act [_]._" [100,100] 100)
    16 
    16 
    17 section {* strong induction principle *}
       
    18 
       
    19 (* FIXME: this induction rule needs some slight change to conform *)
       
    20 (* with the convention from lam_substs                            *)
       
    21 
       
    22 lemma trm_induct_aux:
       
    23   fixes P :: "trm \<Rightarrow> 'a \<Rightarrow> bool"
       
    24   and   f1 :: "'a \<Rightarrow> var set"
       
    25   and   f2 :: "'a \<Rightarrow> mvar set"
       
    26   assumes fs1: "\<And>x. finite (f1 x)"
       
    27       and fs2: "\<And>x. finite (f2 x)"
       
    28       and h1: "\<And>k x. P (Var x) k"  
       
    29       and h2: "\<And>k x t. x\<notin>f1 k \<Longrightarrow> (\<forall>l. P t l) \<Longrightarrow> P (Lam [x].t) k" 
       
    30       and h3: "\<And>k t1 t2. (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l) \<Longrightarrow> P (App t1 t2) k" 
       
    31       and h4: "\<And>k a t1. (\<forall>l. P t1 l) \<Longrightarrow> P (Pss a t1) k" 
       
    32       and h5: "\<And>k a t1. a\<notin>f2 k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> P (Act [a].t1) k"
       
    33   shows "\<forall>(pi1::var prm) (pi2::mvar prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k"
       
    34 proof (induct rule: trm.induct_weak)
       
    35   case (goal1 a)
       
    36   show ?case using h1 by simp
       
    37 next
       
    38   case (goal2 x t)
       
    39   assume i1: "\<forall>(pi1::var prm)(pi2::mvar prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k" 
       
    40   show ?case
       
    41   proof (intro strip, simp add: abs_perm)
       
    42     fix pi1::"var prm" and pi2::"mvar prm" and k::"'a"
       
    43     have f: "\<exists>c::var. c\<sharp>(f1 k,pi1\<bullet>(pi2\<bullet>x),pi1\<bullet>(pi2\<bullet>t))"
       
    44       by (rule at_exists_fresh[OF at_var_inst], simp add: supp_prod fs_var1 
       
    45           at_fin_set_supp[OF at_var_inst, OF fs1] fs1)
       
    46     then obtain c::"var" 
       
    47       where f1: "c\<noteq>(pi1\<bullet>(pi2\<bullet>x))" and f2: "c\<sharp>(f1 k)" and f3: "c\<sharp>(pi1\<bullet>(pi2\<bullet>t))" 
       
    48       by (force simp add: fresh_prod at_fresh[OF at_var_inst])
       
    49     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
       
    50       by (simp add: trm.inject alpha)
       
    51     from i1 have "\<forall>k. P (([(c,pi1\<bullet>(pi2\<bullet>x))]@pi1)\<bullet>(pi2\<bullet>t)) k" by force
       
    52     hence i1b: "\<forall>k. P ([(c,pi1\<bullet>(pi2\<bullet>x))]\<bullet>(pi1\<bullet>(pi2\<bullet>t))) k" by (simp add: pt_var2[symmetric])
       
    53     with h3 f2 have "P (Lam [c].([(c,pi1\<bullet>(pi2\<bullet>x))]\<bullet>(pi1\<bullet>(pi2\<bullet>t)))) k" 
       
    54       by (auto simp add: fresh_def at_fin_set_supp[OF at_var_inst, OF fs1])
       
    55     with g show "P (Lam [(pi1\<bullet>(pi2\<bullet>x))].(pi1\<bullet>(pi2\<bullet>t))) k" by simp 
       
    56   qed
       
    57 next 
       
    58   case (goal3 t1 t2)
       
    59   assume i1: "\<forall>(pi1::var prm)(pi2::mvar prm) k. P (pi1\<bullet>(pi2\<bullet>t1)) k" 
       
    60   assume i2: "\<forall>(pi1::var prm)(pi2::mvar prm) k. P (pi1\<bullet>(pi2\<bullet>t2)) k"
       
    61   show ?case 
       
    62   proof (intro strip)
       
    63     fix pi1::"var prm" and pi2::"mvar prm" and k::"'a"
       
    64     from h3 i1 i2 have "P (App (pi1\<bullet>(pi2\<bullet>t1)) (pi1\<bullet>(pi2\<bullet>t2))) k" by force
       
    65     thus "P (pi1\<bullet>(pi2\<bullet>(App t1 t2))) k" by simp
       
    66   qed
       
    67 next
       
    68   case (goal4 b t)
       
    69   assume i1: "\<forall>(pi1::var prm)(pi2::mvar prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k" 
       
    70   show ?case 
       
    71   proof (intro strip)
       
    72     fix pi1::"var prm" and pi2::"mvar prm" and k::"'a"
       
    73     from h4 i1 have "P (Pss (pi1\<bullet>(pi2\<bullet>b)) (pi1\<bullet>(pi2\<bullet>t))) k" by force
       
    74     thus "P (pi1\<bullet>(pi2\<bullet>(Pss b t))) k" by simp
       
    75   qed
       
    76 next
       
    77   case (goal5 b t)
       
    78   assume i1: "\<forall>(pi1::var prm)(pi2::mvar prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k" 
       
    79   show ?case
       
    80   proof (intro strip, simp add: abs_perm)
       
    81     fix pi1::"var prm" and pi2::"mvar prm" and k::"'a"
       
    82     have f: "\<exists>c::mvar. c\<sharp>(f2 k,pi1\<bullet>(pi2\<bullet>b),pi1\<bullet>(pi2\<bullet>t))"
       
    83       by (rule at_exists_fresh[OF at_mvar_inst], simp add: supp_prod fs_mvar1 
       
    84           at_fin_set_supp[OF at_mvar_inst, OF fs2] fs2)
       
    85     then obtain c::"mvar" 
       
    86       where f1: "c\<noteq>(pi1\<bullet>(pi2\<bullet>b))" and f2: "c\<sharp>(f2 k)" and f3: "c\<sharp>(pi1\<bullet>(pi2\<bullet>t))" 
       
    87       by (force simp add: fresh_prod at_fresh[OF at_mvar_inst])
       
    88     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
       
    89       by (simp add: trm.inject alpha, simp add: dj_cp[OF cp_mvar_var_inst, OF dj_var_mvar])
       
    90     from i1 have "\<forall>k. P (pi1\<bullet>(([(c,pi1\<bullet>(pi2\<bullet>b))]@pi2)\<bullet>t)) k" by force
       
    91     hence i1b: "\<forall>k. P (pi1\<bullet>([(c,pi1\<bullet>(pi2\<bullet>b))]\<bullet>(pi2\<bullet>t))) k" by (simp add: pt_mvar2[symmetric])
       
    92     with h5 f2 have "P (Act [c].(pi1\<bullet>([(c,pi1\<bullet>(pi2\<bullet>b))]\<bullet>(pi2\<bullet>t)))) k" 
       
    93       by (auto simp add: fresh_def at_fin_set_supp[OF at_mvar_inst, OF fs2])
       
    94     with g show "P (Act [(pi1\<bullet>(pi2\<bullet>b))].(pi1\<bullet>(pi2\<bullet>t))) k" by simp 
       
    95   qed
       
    96 qed
       
    97 
       
    98 lemma trm_induct'[case_names Var Lam App Pss Act]:
       
    99   fixes P :: "trm \<Rightarrow> 'a \<Rightarrow> bool"
       
   100   and   f1 :: "'a \<Rightarrow> var set"
       
   101   and   f2 :: "'a \<Rightarrow> mvar set"
       
   102   assumes fs1: "\<And>x. finite (f1 x)"
       
   103       and fs2: "\<And>x. finite (f2 x)"
       
   104       and h1: "\<And>k x. P (Var x) k"  
       
   105       and h2: "\<And>k x t. x\<notin>f1 k \<Longrightarrow> (\<forall>l. P t l) \<Longrightarrow> P (Lam [x].t) k" 
       
   106       and h3: "\<And>k t1 t2. (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l) \<Longrightarrow> P (App t1 t2) k" 
       
   107       and h4: "\<And>k a t1. (\<forall>l. P t1 l) \<Longrightarrow> P (Pss a t1) k" 
       
   108       and h5: "\<And>k a t1. a\<notin>f2 k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> P (Act [a].t1) k"
       
   109   shows "P t k"
       
   110 proof -
       
   111   have "\<forall>(pi1::var prm)(pi2::mvar prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k"
       
   112   using fs1 fs2 h1 h2 h3 h4 h5 by (rule trm_induct_aux, auto)
       
   113   hence "P (([]::var prm)\<bullet>(([]::mvar prm)\<bullet>t)) k" by blast
       
   114   thus "P t k" by simp
       
   115 qed
       
   116 
       
   117 lemma trm_induct[case_names Var Lam App Pss Act]: 
       
   118   fixes P :: "trm \<Rightarrow> ('a::{fs_var,fs_mvar}) \<Rightarrow> bool"
       
   119   assumes h1: "\<And>k x. P (Var x) k"  
       
   120       and h2: "\<And>k x t. x\<sharp>k \<Longrightarrow> (\<forall>l. P t l) \<Longrightarrow> P (Lam [x].t) k" 
       
   121       and h3: "\<And>k t1 t2. (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l) \<Longrightarrow> P (App t1 t2) k" 
       
   122       and h4: "\<And>k a t1. (\<forall>l. P t1 l) \<Longrightarrow> P (Pss a t1) k" 
       
   123       and h5: "\<And>k a t1. a\<sharp>k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> P (Act [a].t1) k"
       
   124   shows  "P t k"
       
   125 by (rule trm_induct'[of "\<lambda>x. ((supp x)::var set)" "\<lambda>x. ((supp x)::mvar set)" "P"], 
       
   126     simp_all add: fs_var1 fs_mvar1 fresh_def[symmetric], auto intro: h1 h2 h3 h4 h5)