src/HOL/Nominal/Examples/Class.thy
author urbanc
Fri, 16 Dec 2005 18:20:59 +0100
changeset 18425 bcf13dbaa339
parent 18395 87217764cec2
child 18661 dde117622dac
permissions -rw-r--r--
I think the earlier version was completely broken (not sure about this one)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18425
bcf13dbaa339 I think the earlier version was completely broken
urbanc
parents: 18395
diff changeset
     1
bcf13dbaa339 I think the earlier version was completely broken
urbanc
parents: 18395
diff changeset
     2
theory class 
bcf13dbaa339 I think the earlier version was completely broken
urbanc
parents: 18395
diff changeset
     3
imports "../nominal" 
bcf13dbaa339 I think the earlier version was completely broken
urbanc
parents: 18395
diff changeset
     4
begin
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
     5
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
     6
atom_decl name coname
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
     7
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
     8
section {* Term-Calculus from my PHD *}
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
     9
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    10
nominal_datatype trm = Ax  "name" "coname"
18425
bcf13dbaa339 I think the earlier version was completely broken
urbanc
parents: 18395
diff changeset
    11
                 | ImpR "\<guillemotleft>name\<guillemotright>(\<guillemotleft>coname\<guillemotright>trm)" "coname"  ("ImpR [_].[_]._ _" [100,100,100,100] 100)
bcf13dbaa339 I think the earlier version was completely broken
urbanc
parents: 18395
diff changeset
    12
                 | ImpL "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"("ImpL [_]._ [_]._ _" [100,100,100,100,100] 100)
bcf13dbaa339 I think the earlier version was completely broken
urbanc
parents: 18395
diff changeset
    13
                 | Cut "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm"        ("Cut [_]._ [_]._" [100,100,100,100] 100)
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    14
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    15
lemma trm_induct_aux:
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    16
  fixes P :: "trm \<Rightarrow> 'a \<Rightarrow> bool"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    17
  and   f1 :: "'a \<Rightarrow> name set"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    18
  and   f2 :: "'a \<Rightarrow> coname set"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    19
  assumes fs1: "\<And>x. finite (f1 x)"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    20
      and fs2: "\<And>x. finite (f2 x)"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    21
      and h1: "\<And>k x a. P (Ax x a) k"  
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    22
      and h2: "\<And>k x a t b. x\<notin>f1 k \<Longrightarrow> a\<notin>f2 k \<Longrightarrow> (\<forall>l. P t l) \<Longrightarrow> P (ImpR [x].[a].t b) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    23
      and h3: "\<And>k a t1 x t2 y. a\<notin>f2 k \<Longrightarrow> x\<notin>f1 k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l) 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    24
               \<Longrightarrow> P (ImpL [a].t1 [x].t2 y) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    25
      and h4: "\<And>k a t1 x t2. a\<notin>f2 k \<Longrightarrow> x\<notin>f1 k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l) 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    26
               \<Longrightarrow> P (Cut [a].t1 [x].t2) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    27
  shows "\<forall>(pi1::name prm) (pi2::coname prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k"
18425
bcf13dbaa339 I think the earlier version was completely broken
urbanc
parents: 18395
diff changeset
    28
proof (induct rule: trm.induct_weak)
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    29
  case (goal1 a)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    30
  show ?case using h1 by simp
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    31
next
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    32
  case (goal2 x a t b)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    33
  assume i1: "\<forall>(pi1::name prm)(pi2::coname prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    34
  show ?case 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    35
  proof (intro strip, simp add: abs_perm perm_dj)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    36
    fix pi1::"name prm" and pi2::"coname prm" and k::"'a"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    37
    have "\<exists>u::name. u\<sharp>(f1 k,pi1\<bullet>x,pi1\<bullet>(pi2\<bullet>t))"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    38
      by (rule at_exists_fresh[OF at_name_inst], simp add: supp_prod fs_name1 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    39
               at_fin_set_supp[OF at_name_inst, OF fs1] fs1)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    40
    then obtain u::"name" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    41
      where f1: "u\<noteq>(pi1\<bullet>x)" and f2: "u\<sharp>(f1 k)" and f3: "u\<sharp>(pi1\<bullet>(pi2\<bullet>t))" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    42
      by (auto simp add: fresh_prod at_fresh[OF at_name_inst])
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    43
    have "\<exists>c::coname. c\<sharp>(f2 k,pi2\<bullet>a,pi1\<bullet>(pi2\<bullet>t))"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    44
      by (rule at_exists_fresh[OF at_coname_inst], simp add: supp_prod fs_coname1 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    45
               at_fin_set_supp[OF at_coname_inst, OF fs2] fs2)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    46
    then obtain c::"coname" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    47
      where e1: "c\<noteq>(pi2\<bullet>a)" and e2: "c\<sharp>(f2 k)" and e3: "c\<sharp>(pi1\<bullet>(pi2\<bullet>t))" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    48
      by (auto simp add: fresh_prod at_fresh[OF at_coname_inst])
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    49
    have g: "ImpR [u].[c].([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>([(c,pi2\<bullet>a)]\<bullet>(pi2\<bullet>t)))) (pi2\<bullet>b)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    50
            =ImpR [(pi1\<bullet>x)].[(pi2\<bullet>a)].(pi1\<bullet>(pi2\<bullet>t)) (pi2\<bullet>b)" using f1 f3 e1 e3
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    51
      apply (auto simp add: ImpR_inject alpha abs_fresh abs_perm perm_dj,
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    52
                  simp add: dj_cp[OF cp_name_coname_inst, OF dj_coname_name])
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    53
      apply(simp add:  pt_fresh_left_ineq[OF pt_name_inst, OF pt_name_inst, 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    54
                                          OF at_name_inst, OF cp_name_coname_inst] perm_dj)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    55
      done
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    56
    from i1 have "\<forall>k. P (([(u,pi1\<bullet>x)]@pi1)\<bullet>(([(c,pi2\<bullet>a)]@pi2)\<bullet>t)) k" by force
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    57
    hence i1b: "\<forall>k. P ([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>([(c,pi2\<bullet>a)]\<bullet>(pi2\<bullet>t)))) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    58
      by (simp add: pt_name2[symmetric] pt_coname2[symmetric])
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    59
    with h2 f2 e2 have "P (ImpR [u].[c].([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>([(c,pi2\<bullet>a)]\<bullet>(pi2\<bullet>t)))) (pi2\<bullet>b)) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    60
      by (simp add: fresh_def at_fin_set_supp[OF at_name_inst, OF fs1]
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    61
                                   at_fin_set_supp[OF at_coname_inst, OF fs2])
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    62
    with g show "P (ImpR [(pi1\<bullet>x)].[(pi2\<bullet>a)].(pi1\<bullet>(pi2\<bullet>t)) (pi2\<bullet>b)) k" by simp 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    63
  qed
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    64
next
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    65
  case (goal3 a t1 x t2 y)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    66
  assume i1: "\<forall>(pi1::name prm)(pi2::coname prm) k. P (pi1\<bullet>(pi2\<bullet>t1)) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    67
  and    i2: "\<forall>(pi1::name prm)(pi2::coname prm) k. P (pi1\<bullet>(pi2\<bullet>t2)) k"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    68
  show ?case
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    69
  proof (intro strip, simp add: abs_perm)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    70
    fix pi1::"name prm" and pi2::"coname prm" and k::"'a"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    71
    have "\<exists>u::name. u\<sharp>(f1 k,pi1\<bullet>x,pi1\<bullet>(pi2\<bullet>t2))"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    72
      by (rule at_exists_fresh[OF at_name_inst], simp add: supp_prod fs_name1 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    73
               at_fin_set_supp[OF at_name_inst, OF fs1] fs1)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    74
    then obtain u::"name" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    75
      where f1: "u\<noteq>(pi1\<bullet>x)" and f2: "u\<sharp>(f1 k)" and f3: "u\<sharp>(pi1\<bullet>(pi2\<bullet>t2))" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    76
      by (auto simp add: fresh_prod at_fresh[OF at_name_inst])
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    77
    have "\<exists>c::coname. c\<sharp>(f2 k,pi2\<bullet>a,pi1\<bullet>(pi2\<bullet>t1))"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    78
      by (rule at_exists_fresh[OF at_coname_inst], simp add: supp_prod fs_coname1 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    79
               at_fin_set_supp[OF at_coname_inst, OF fs2] fs2)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    80
    then obtain c::"coname" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    81
      where e1: "c\<noteq>(pi2\<bullet>a)" and e2: "c\<sharp>(f2 k)" and e3: "c\<sharp>(pi1\<bullet>(pi2\<bullet>t1))" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    82
      by (auto simp add: fresh_prod at_fresh[OF at_coname_inst])
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    83
    have g: "ImpL [c].([(c,pi2\<bullet>a)]\<bullet>(pi1\<bullet>(pi2\<bullet>t1))) [u].([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>(pi2\<bullet>t2))) (pi1\<bullet>y)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    84
            =ImpL [(pi2\<bullet>a)].(pi1\<bullet>(pi2\<bullet>t1)) [(pi1\<bullet>x)].(pi1\<bullet>(pi2\<bullet>t2)) (pi1\<bullet>y)" using f1 f3 e1 e3
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    85
      by (simp add: ImpL_inject alpha abs_fresh abs_perm perm_dj)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    86
    from i2 have "\<forall>k. P (([(u,pi1\<bullet>x)]@pi1)\<bullet>(pi2\<bullet>t2)) k" by force
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    87
    hence i2b: "\<forall>k. P ([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>(pi2\<bullet>t2))) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    88
      by (simp add: pt_name2[symmetric])
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    89
    from i1 have "\<forall>k. P (pi1\<bullet>(([(c,pi2\<bullet>a)]@pi2)\<bullet>t1)) k" by force
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    90
    hence i1b: "\<forall>k. P ([(c,pi2\<bullet>a)]\<bullet>(pi1\<bullet>(pi2\<bullet>t1))) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    91
      by (simp add: pt_coname2[symmetric] dj_cp[OF cp_name_coname_inst, OF dj_coname_name])
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    92
    from h3 f2 e2 i1b i2b 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    93
    have "P (ImpL [c].([(c,pi2\<bullet>a)]\<bullet>(pi1\<bullet>(pi2\<bullet>t1))) [u].([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>(pi2\<bullet>t2))) (pi1\<bullet>y)) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    94
      by (simp add: fresh_def at_fin_set_supp[OF at_name_inst, OF fs1]
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    95
                                   at_fin_set_supp[OF at_coname_inst, OF fs2])
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    96
    with g show "P (ImpL [(pi2\<bullet>a)].(pi1\<bullet>(pi2\<bullet>t1)) [(pi1\<bullet>x)].(pi1\<bullet>(pi2\<bullet>t2)) (pi1\<bullet>y)) k" by simp 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    97
  qed
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    98
next
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    99
  case (goal4 a t1 x t2)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   100
  assume i1: "\<forall>(pi1::name prm)(pi2::coname prm) k. P (pi1\<bullet>(pi2\<bullet>t1)) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   101
  and    i2: "\<forall>(pi1::name prm)(pi2::coname prm) k. P (pi1\<bullet>(pi2\<bullet>t2)) k"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   102
  show ?case
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   103
  proof (intro strip, simp add: abs_perm)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   104
    fix pi1::"name prm" and pi2::"coname prm" and k::"'a"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   105
    have "\<exists>u::name. u\<sharp>(f1 k,pi1\<bullet>x,pi1\<bullet>(pi2\<bullet>t2))"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   106
      by (rule at_exists_fresh[OF at_name_inst], simp add: supp_prod fs_name1 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   107
               at_fin_set_supp[OF at_name_inst, OF fs1] fs1)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   108
    then obtain u::"name" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   109
      where f1: "u\<noteq>(pi1\<bullet>x)" and f2: "u\<sharp>(f1 k)" and f3: "u\<sharp>(pi1\<bullet>(pi2\<bullet>t2))" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   110
      by (auto simp add: fresh_prod at_fresh[OF at_name_inst])
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   111
    have "\<exists>c::coname. c\<sharp>(f2 k,pi2\<bullet>a,pi1\<bullet>(pi2\<bullet>t1))"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   112
      by (rule at_exists_fresh[OF at_coname_inst], simp add: supp_prod fs_coname1 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   113
               at_fin_set_supp[OF at_coname_inst, OF fs2] fs2)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   114
    then obtain c::"coname" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   115
      where e1: "c\<noteq>(pi2\<bullet>a)" and e2: "c\<sharp>(f2 k)" and e3: "c\<sharp>(pi1\<bullet>(pi2\<bullet>t1))" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   116
      by (auto simp add: fresh_prod at_fresh[OF at_coname_inst])
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   117
    have g: "Cut [c].([(c,pi2\<bullet>a)]\<bullet>(pi1\<bullet>(pi2\<bullet>t1))) [u].([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>(pi2\<bullet>t2)))
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   118
            =Cut [(pi2\<bullet>a)].(pi1\<bullet>(pi2\<bullet>t1)) [(pi1\<bullet>x)].(pi1\<bullet>(pi2\<bullet>t2))" using f1 f3 e1 e3
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   119
      by (simp add: Cut_inject alpha abs_fresh abs_perm perm_dj)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   120
    from i2 have "\<forall>k. P (([(u,pi1\<bullet>x)]@pi1)\<bullet>(pi2\<bullet>t2)) k" by force
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   121
    hence i2b: "\<forall>k. P ([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>(pi2\<bullet>t2))) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   122
      by (simp add: pt_name2[symmetric])
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   123
    from i1 have "\<forall>k. P (pi1\<bullet>(([(c,pi2\<bullet>a)]@pi2)\<bullet>t1)) k" by force
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   124
    hence i1b: "\<forall>k. P ([(c,pi2\<bullet>a)]\<bullet>(pi1\<bullet>(pi2\<bullet>t1))) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   125
      by (simp add: pt_coname2[symmetric] dj_cp[OF cp_name_coname_inst, OF dj_coname_name])
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   126
    from h3 f2 e2 i1b i2b 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   127
    have "P (Cut [c].([(c,pi2\<bullet>a)]\<bullet>(pi1\<bullet>(pi2\<bullet>t1))) [u].([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>(pi2\<bullet>t2)))) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   128
      by (simp add: fresh_def at_fin_set_supp[OF at_name_inst, OF fs1]
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   129
                                   at_fin_set_supp[OF at_coname_inst, OF fs2])
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   130
    with g show "P (Cut [(pi2\<bullet>a)].(pi1\<bullet>(pi2\<bullet>t1)) [(pi1\<bullet>x)].(pi1\<bullet>(pi2\<bullet>t2))) k" by simp 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   131
  qed
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   132
qed
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   133
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   134
lemma trm_induct'[case_names Ax ImpR ImpL Cut]: 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   135
  fixes P :: "trm \<Rightarrow> 'a \<Rightarrow> bool"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   136
  and   f1 :: "'a \<Rightarrow> name set"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   137
  and   f2 :: "'a \<Rightarrow> coname set"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   138
  assumes fs1: "\<And>x. finite (f1 x)"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   139
      and fs2: "\<And>x. finite (f2 x)"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   140
      and h1: "\<And>k x a. P (Ax x a) k"  
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   141
      and h2: "\<And>k x a t b. x\<notin>f1 k \<Longrightarrow> a\<notin>f2 k \<Longrightarrow> (\<forall>l. P t l) \<Longrightarrow> P (ImpR [x].[a].t b) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   142
      and h3: "\<And>k a t1 x t2 y. a\<notin>f2 k \<Longrightarrow> x\<notin>f1 k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l) 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   143
               \<Longrightarrow> P (ImpL [a].t1 [x].t2 y) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   144
      and h4: "\<And>k a t1 x t2. a\<notin>f2 k \<Longrightarrow> x\<notin>f1 k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l) 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   145
               \<Longrightarrow> P (Cut [a].t1 [x].t2) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   146
  shows  "P t k"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   147
proof -
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   148
  have "\<forall>(pi1::name prm)(pi2::coname prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   149
  using fs1 fs2 h1 h2 h3 h4 by (rule trm_induct_aux, auto)
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   150
  hence "P (([]::name prm)\<bullet>(([]::coname prm)\<bullet>t)) k" by blast
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   151
  thus "P t k" by simp
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   152
qed
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   153
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   154
lemma trm_induct[case_names Ax ImpR ImpL Cut]: 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   155
  fixes P :: "trm \<Rightarrow> ('a::{fs_name,fs_coname}) \<Rightarrow> bool"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   156
  assumes h1: "\<And>k x a. P (Ax x a) k"  
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   157
      and h2: "\<And>k x a t b. x\<sharp>k \<Longrightarrow> a\<sharp>k \<Longrightarrow> (\<forall>l. P t l) \<Longrightarrow> P (ImpR [x].[a].t b) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   158
      and h3: "\<And>k a t1 x t2 y. a\<sharp>k \<Longrightarrow> x\<sharp>k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l) 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   159
               \<Longrightarrow> P (ImpL [a].t1 [x].t2 y) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   160
      and h4: "\<And>k a t1 x t2. a\<sharp>k \<Longrightarrow> x\<sharp>k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l) 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   161
               \<Longrightarrow> P (Cut [a].t1 [x].t2) k" 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   162
  shows  "P t k"
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   163
by (rule trm_induct'[of "\<lambda>x. ((supp x)::name set)" "\<lambda>x. ((supp x)::coname set)" "P"], 
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
   164
    simp_all add: fs_name1 fs_coname1 fresh_def[symmetric], auto intro: h1 h2 h3 h4)