src/HOL/Nominal/Examples/Class.thy
author berghofe
Wed, 07 Feb 2007 17:44:07 +0100
changeset 22271 51a80e238b29
parent 22232 340cb955008e
child 22418 49e2d9744ae1
permissions -rw-r--r--
Adapted to new inductive definition package.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19500
188d4e44c1a6 Removed broken proof.
berghofe
parents: 19477
diff changeset
     1
(* $Id$ *)
18425
bcf13dbaa339 I think the earlier version was completely broken
urbanc
parents: 18395
diff changeset
     2
18661
dde117622dac updated to new induction principle
urbanc
parents: 18425
diff changeset
     3
theory Class 
22232
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
     4
imports "Nominal" 
18425
bcf13dbaa339 I think the earlier version was completely broken
urbanc
parents: 18395
diff changeset
     5
begin
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
     6
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
     7
section {* Term-Calculus from Urban's PhD *}
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
     8
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
     9
atom_decl name coname
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    10
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    11
nominal_datatype trm = 
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    12
    Ax   "name" "coname"
19319
7e1f85ceb1a2 added the first two simple proofs of the recursion
urbanc
parents: 18881
diff changeset
    13
  | Cut  "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm"            ("Cut <_>._ [_]._" [100,100,100,100] 100)
7e1f85ceb1a2 added the first two simple proofs of the recursion
urbanc
parents: 18881
diff changeset
    14
  | NotR "\<guillemotleft>name\<guillemotright>trm" "coname"                 ("NotR [_]._ _" [100,100,100] 100)
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    15
  | NotL "\<guillemotleft>coname\<guillemotright>trm" "name"                 ("NotL <_>._ _" [100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    16
  | AndR "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>coname\<guillemotright>trm" "coname" ("AndR <_>._ <_>._ _" [100,100,100,100,100] 100)
19319
7e1f85ceb1a2 added the first two simple proofs of the recursion
urbanc
parents: 18881
diff changeset
    17
  | AndL1 "\<guillemotleft>name\<guillemotright>trm" "name"                  ("AndL1 [_]._ _" [100,100,100] 100)
7e1f85ceb1a2 added the first two simple proofs of the recursion
urbanc
parents: 18881
diff changeset
    18
  | AndL2 "\<guillemotleft>name\<guillemotright>trm" "name"                  ("AndL2 [_]._ _" [100,100,100] 100)
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    19
  | OrR1 "\<guillemotleft>coname\<guillemotright>trm" "coname"               ("OrR1 <_>._ _" [100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    20
  | OrR2 "\<guillemotleft>coname\<guillemotright>trm" "coname"               ("OrR2 <_>._ _" [100,100,100] 100)
19319
7e1f85ceb1a2 added the first two simple proofs of the recursion
urbanc
parents: 18881
diff changeset
    21
  | OrL "\<guillemotleft>name\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"        ("OrL [_]._ [_]._ _" [100,100,100,100,100] 100)
7e1f85ceb1a2 added the first two simple proofs of the recursion
urbanc
parents: 18881
diff changeset
    22
  | ImpR "\<guillemotleft>name\<guillemotright>(\<guillemotleft>coname\<guillemotright>trm)" "coname"       ("ImpR [_].<_>._ _" [100,100,100,100] 100)
7e1f85ceb1a2 added the first two simple proofs of the recursion
urbanc
parents: 18881
diff changeset
    23
  | ImpL "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"     ("ImpL <_>._ [_]._ _" [100,100,100,100,100] 100)
7e1f85ceb1a2 added the first two simple proofs of the recursion
urbanc
parents: 18881
diff changeset
    24
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    25
text {* Induction principles *}
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    26
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    27
thm trm.induct_weak --"weak"
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    28
thm trm.induct      --"strong"
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    29
thm trm.induct'     --"strong with explicit context (rarely needed)"
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    30
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    31
text {* named terms *}
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    32
22232
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    33
nominal_datatype ntrm = Na "\<guillemotleft>name\<guillemotright>trm" ("([_]:_)" [100,100] 100)
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    34
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    35
text {* conamed terms *}
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    36
22232
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    37
nominal_datatype ctrm = Co "\<guillemotleft>coname\<guillemotright>trm" ("(<_>:_)" [100,100] 100)
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    38
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    39
lemma eq_eqvt_name[eqvt]:
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    40
  fixes pi::"name prm"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    41
  and   x::"'a::pt_name"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    42
  shows "pi\<bullet>(x=y) = ((pi\<bullet>x)=(pi\<bullet>y))"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    43
  by (simp add: perm_bool perm_bij)
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    44
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    45
lemma eq_eqvt_coname[eqvt]:
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    46
  fixes pi::"coname prm"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    47
  and   x::"'a::pt_coname"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    48
  shows "pi\<bullet>(x=y) = ((pi\<bullet>x)=(pi\<bullet>y))"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    49
  by (simp add: perm_bool perm_bij)
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    50
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    51
text {* renaming functions *}
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    52
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    53
consts
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    54
  nrename :: "trm \<Rightarrow> name \<Rightarrow> name \<Rightarrow> trm"      ("_[_\<turnstile>n>_]" [100,100,100] 100) 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    55
  crename :: "trm \<Rightarrow> coname \<Rightarrow> coname \<Rightarrow> trm"  ("_[_\<turnstile>c>_]" [100,100,100] 100) 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    56
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    57
nominal_primrec (freshness_context: "(d::coname,e::coname)") 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    58
  "(Ax x a)[d\<turnstile>c>e] = (if a=d then Ax x e else Ax x a)" 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    59
  "\<lbrakk>a\<sharp>(d,e,N);x\<sharp>M\<rbrakk> \<Longrightarrow> (Cut <a>.M [x].N)[d\<turnstile>c>e] = Cut <a>.(M[d\<turnstile>c>e]) [x].(N[d\<turnstile>c>e])" 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    60
  "(NotR [x].M a)[d\<turnstile>c>e] = (if a=d then NotR [x].(M[d\<turnstile>c>e]) e else NotR [x].(M[d\<turnstile>c>e]) a)" 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    61
  "a\<sharp>(d,e) \<Longrightarrow> (NotL <a>.M x)[d\<turnstile>c>e] = (NotL <a>.(M[d\<turnstile>c>e]) x)" 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    62
  "\<lbrakk>a\<sharp>(d,e,N,c);b\<sharp>(d,e,M,c);b\<noteq>a\<rbrakk> \<Longrightarrow> (AndR <a>.M <b>.N c)[d\<turnstile>c>e] = 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    63
          (if c=d then AndR <a>.(M[d\<turnstile>c>e]) <b>.(N[d \<turnstile>c>e]) e else AndR <a>.(M[d\<turnstile>c>e]) <b>.(N[d\<turnstile>c>e]) c)" 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    64
  "x\<sharp>y \<Longrightarrow> (AndL1 [x].M y)[d\<turnstile>c>e] = AndL1 [x].(M[d\<turnstile>c>e]) y"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    65
  "x\<sharp>y \<Longrightarrow> (AndL2 [x].M y)[d\<turnstile>c>e] = AndL2 [x].(M[d\<turnstile>c>e]) y"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    66
  "a\<sharp>(d,e,b) \<Longrightarrow> (OrR1 <a>.M b)[d\<turnstile>c>e] = (if b=d then OrR1 <a>.(M[d\<turnstile>c>e]) e else OrR1 <a>.(M[d\<turnstile>c>e]) b)"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    67
  "a\<sharp>(d,e,b) \<Longrightarrow> (OrR2 <a>.M b)[d\<turnstile>c>e] = (if b=d then OrR2 <a>.(M[d\<turnstile>c>e]) e else OrR2 <a>.(M[d\<turnstile>c>e]) b)"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    68
  "\<lbrakk>x\<sharp>(N,z);y\<sharp>(M,z);y\<noteq>x\<rbrakk> \<Longrightarrow> (OrL [x].M [y].N z)[d\<turnstile>c>e] = OrL [x].(M[d\<turnstile>c>e]) [y].(N[d\<turnstile>c>e]) z"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    69
  "a\<sharp>(d,e,b) \<Longrightarrow> (ImpR [x].<a>.M b)[d\<turnstile>c>e] = 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    70
       (if b=d then ImpR [x].<a>.(M[d\<turnstile>c>e]) e else ImpR [x].<a>.(M[d\<turnstile>c>e]) b)"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    71
  "\<lbrakk>a\<sharp>(d,e,N);x\<sharp>(M,y)\<rbrakk> \<Longrightarrow> (ImpL <a>.M [x].N y)[d\<turnstile>c>e] = ImpL <a>.(M[d\<turnstile>c>e]) [x].(N[d\<turnstile>c>e]) y"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    72
apply(finite_guess add: eqvt perm_if fs_coname1 fs_name1 | 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    73
      perm_simp add: abs_fresh abs_supp fs_name1 fs_coname1)+
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    74
apply(fresh_guess add: eqvt perm_if fs_coname1 fs_name1 | perm_simp add: fresh_atm)+
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    75
done
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    76
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    77
nominal_primrec (freshness_context: "(u::name,v::name)") 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    78
  "(Ax x a)[u\<turnstile>n>v] = (if x=u then Ax v a else Ax x a)" 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    79
  "\<lbrakk>a\<sharp>N;x\<sharp>(u,v,M)\<rbrakk> \<Longrightarrow> (Cut <a>.M [x].N)[u\<turnstile>n>v] = Cut <a>.(M[u\<turnstile>n>v]) [x].(N[u\<turnstile>n>v])" 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    80
  "x\<sharp>(u,v) \<Longrightarrow> (NotR [x].M a)[u\<turnstile>n>v] = NotR [x].(M[u\<turnstile>n>v]) a" 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    81
  "(NotL <a>.M x)[u\<turnstile>n>v] = (if x=u then NotL <a>.(M[u\<turnstile>n>v]) v else NotL <a>.(M[u\<turnstile>n>v]) x)" 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    82
  "\<lbrakk>a\<sharp>(N,c);b\<sharp>(M,c);b\<noteq>a\<rbrakk> \<Longrightarrow> (AndR <a>.M <b>.N c)[u\<turnstile>n>v] = AndR <a>.(M[u\<turnstile>n>v]) <b>.(N[u\<turnstile>n>v]) c" 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    83
  "x\<sharp>(u,v,y) \<Longrightarrow> (AndL1 [x].M y)[u\<turnstile>n>v] = (if y=u then AndL1 [x].(M[u\<turnstile>n>v]) v else AndL1 [x].(M[u\<turnstile>n>v]) y)"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    84
  "x\<sharp>(u,v,y) \<Longrightarrow> (AndL2 [x].M y)[u\<turnstile>n>v] = (if y=u then AndL2 [x].(M[u\<turnstile>n>v]) v else AndL2 [x].(M[u\<turnstile>n>v]) y)"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    85
  "a\<sharp>b \<Longrightarrow> (OrR1 <a>.M b)[u\<turnstile>n>v] = OrR1 <a>.(M[u\<turnstile>n>v]) b"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    86
  "a\<sharp>b \<Longrightarrow> (OrR2 <a>.M b)[u\<turnstile>n>v] = OrR2 <a>.(M[u\<turnstile>n>v]) b"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    87
  "\<lbrakk>x\<sharp>(u,v,N,z);y\<sharp>(u,v,M,z);y\<noteq>x\<rbrakk> \<Longrightarrow> (OrL [x].M [y].N z)[u\<turnstile>n>v] = 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    88
        (if z=u then OrL [x].(M[u\<turnstile>n>v]) [y].(N[u\<turnstile>n>v]) v else OrL [x].(M[u\<turnstile>n>v]) [y].(N[u\<turnstile>n>v]) z)"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    89
  "\<lbrakk>a\<sharp>b; x\<sharp>(u,v)\<rbrakk> \<Longrightarrow> (ImpR [x].<a>.M b)[u\<turnstile>n>v] = ImpR [x].<a>.(M[u\<turnstile>n>v]) b"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    90
  "\<lbrakk>a\<sharp>N;x\<sharp>(u,v,M,y)\<rbrakk> \<Longrightarrow> (ImpL <a>.M [x].N y)[u\<turnstile>n>v] = 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    91
        (if y=u then ImpL <a>.(M[u\<turnstile>n>v]) [x].(N[u\<turnstile>n>v]) v else ImpL <a>.(M[u\<turnstile>n>v]) [x].(N[u\<turnstile>n>v]) y)"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    92
apply(finite_guess add: eqvt perm_if fs_coname1 fs_name1 | 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    93
      perm_simp add: abs_fresh abs_supp fresh_prod fs_name1 fs_coname1)+
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    94
apply(fresh_guess add: eqvt perm_if fs_coname1 fs_name1 | perm_simp add: fresh_atm)+
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    95
done
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    96
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    97
text {* We should now define the two forms of substitition :o( *}
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    98
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    99
consts
22232
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
   100
  substn :: "trm \<Rightarrow> name   \<Rightarrow> ctrm \<Rightarrow> trm" ("_[_::n=_]" [100,100,100] 100) 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
   101
  substc :: "trm \<Rightarrow> coname \<Rightarrow> ntrm \<Rightarrow> trm" ("_[_::c=_]" [100,100,100] 100)
19500
188d4e44c1a6 Removed broken proof.
berghofe
parents: 19477
diff changeset
   102
188d4e44c1a6 Removed broken proof.
berghofe
parents: 19477
diff changeset
   103
end
188d4e44c1a6 Removed broken proof.
berghofe
parents: 19477
diff changeset
   104