src/HOL/Nominal/Examples/Compile.thy
author wenzelm
Sun, 10 Dec 2006 15:30:45 +0100
changeset 21746 9d0652354513
parent 21543 e855f25df0c8
child 22271 51a80e238b29
permissions -rw-r--r--
LocalTheory.notation/abbrev;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
21087
3e56528a39f7 Adapted to changes in FCBs.
berghofe
parents: 20955
diff changeset
     1
(* $Id$ *)
20606
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
     2
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
     3
(* A challenge suggested by Adam Chlipala *)
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
     4
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
     5
theory Compile
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
     6
imports "Nominal"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
     7
begin
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
     8
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
     9
atom_decl name 
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    10
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    11
nominal_datatype data = DNat
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    12
                      | DProd "data" "data"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    13
                      | DSum "data" "data"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    14
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    15
nominal_datatype ty = Data "data"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    16
                    | Arrow "ty" "ty" ("_\<rightarrow>_" [100,100] 100)
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    17
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    18
nominal_datatype trm = Var "name"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    19
                     | Lam "\<guillemotleft>name\<guillemotright>trm" ("Lam [_]._" [100,100] 100)
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    20
                     | App "trm" "trm"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    21
                     | Const "nat"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    22
                     | Pr "trm" "trm"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    23
                     | Fst "trm"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    24
                     | Snd "trm"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    25
                     | InL "trm"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    26
                     | InR "trm"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    27
                     | Case "trm" "\<guillemotleft>name\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" 
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    28
                             ("Case _ of inl _ \<rightarrow> _ | inr _ \<rightarrow> _" [100,100,100,100,100] 100)
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    29
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    30
nominal_datatype dataI = OneI | NatI
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    31
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    32
nominal_datatype tyI = DataI "dataI"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    33
                     | ArrowI "tyI" "tyI" ("_\<rightarrow>_" [100,100] 100)
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    34
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    35
nominal_datatype trmI = IVar "name"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    36
                      | ILam "\<guillemotleft>name\<guillemotright>trmI" ("ILam [_]._" [100,100] 100)
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    37
                      | IApp "trmI" "trmI"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    38
                      | IUnit
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    39
                      | INat "nat"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    40
                      | ISucc "trmI"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    41
                      | IAss "trmI" "trmI" ("_\<mapsto>_" [100,100] 100)
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    42
                      | IRef "trmI" 
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    43
                      | ISeq "trmI" "trmI" ("_;;_" [100,100] 100)
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    44
                      | Iif "trmI" "trmI" "trmI"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    45
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    46
text {* valid contexts *}
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    47
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    48
consts
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    49
  ctxts :: "((name\<times>'a::pt_name) list) set" 
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    50
  valid :: "(name\<times>'a::pt_name) list \<Rightarrow> bool"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    51
translations
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    52
  "valid \<Gamma>" \<rightleftharpoons> "\<Gamma> \<in> ctxts"  
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    53
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    54
inductive ctxts
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    55
intros
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    56
v1[intro]: "valid []"
20955
65a9a30b8ece made some proof look more like the ones in Barendregt
urbanc
parents: 20606
diff changeset
    57
v2[intro]: "\<lbrakk>valid \<Gamma>;a\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((a,\<sigma>)#\<Gamma>)" (* maybe dom of \<Gamma> *)
20606
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    58
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    59
text {* typing judgements for trms *}
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    60
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    61
consts
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    62
  typing :: "(((name\<times>ty) list)\<times>trm\<times>ty) set" 
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    63
syntax
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    64
  "_typing_judge" :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [80,80,80] 80) 
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    65
translations
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    66
  "\<Gamma> \<turnstile> t : \<tau>" \<rightleftharpoons> "(\<Gamma>,t,\<tau>) \<in> typing"  
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    67
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    68
inductive typing
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    69
intros
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    70
t0[intro]: "\<lbrakk>valid \<Gamma>; (x,\<tau>)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var x : \<tau>"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    71
t1[intro]: "\<lbrakk>\<Gamma> \<turnstile> e1 : \<tau>1\<rightarrow>\<tau>2; \<Gamma> \<turnstile> e2 : \<tau>1\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App e1 e2 : \<tau>2"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    72
t2[intro]: "\<lbrakk>x\<sharp>\<Gamma>;((x,\<tau>1)#\<Gamma>) \<turnstile> t : \<tau>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : \<tau>1\<rightarrow>\<tau>2"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    73
t3[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Const n : Data(DNat)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    74
t4[intro]: "\<lbrakk>\<Gamma> \<turnstile> e1 : Data(\<sigma>1); \<Gamma> \<turnstile> e2 : Data(\<sigma>2)\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Pr e1 e2 : Data (DProd \<sigma>1 \<sigma>2)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    75
t5[intro]: "\<lbrakk>\<Gamma> \<turnstile> e : Data(DProd \<sigma>1 \<sigma>2)\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Fst e : Data(\<sigma>1)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    76
t6[intro]: "\<lbrakk>\<Gamma> \<turnstile> e : Data(DProd \<sigma>1 \<sigma>2)\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Snd e : Data(\<sigma>2)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    77
t7[intro]: "\<lbrakk>\<Gamma> \<turnstile> e : Data(\<sigma>1)\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> InL e : Data(DSum \<sigma>1 \<sigma>2)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    78
t8[intro]: "\<lbrakk>\<Gamma> \<turnstile> e : Data(\<sigma>2)\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> InR e : Data(DSum \<sigma>1 \<sigma>2)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    79
t9[intro]: "\<lbrakk>x1\<sharp>\<Gamma>; x2\<sharp>\<Gamma>; \<Gamma> \<turnstile> e: Data(DSum \<sigma>1 \<sigma>2); 
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    80
             ((x1,Data(\<sigma>1))#\<Gamma>) \<turnstile> e1 : \<tau>; ((x2,Data(\<sigma>2))#\<Gamma>) \<turnstile> e2 : \<tau>\<rbrakk> 
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    81
             \<Longrightarrow> \<Gamma> \<turnstile> (Case e of inl x1 \<rightarrow> e1 | inr x2 \<rightarrow> e2) : \<tau>"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    82
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    83
text {* typing judgements for Itrms *}
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    84
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    85
consts
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    86
  Ityping :: "(((name\<times>tyI) list)\<times>trmI\<times>tyI) set" 
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    87
syntax
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    88
  "_typing_judge" :: "(name\<times>tyI) list\<Rightarrow>trmI\<Rightarrow>tyI\<Rightarrow>bool" (" _ I\<turnstile> _ : _ " [80,80,80] 80) 
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    89
translations
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    90
  "\<Gamma> I\<turnstile> t : \<tau>" \<rightleftharpoons> "(\<Gamma>,t,\<tau>) \<in> Ityping"  
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    91
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    92
inductive Ityping
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    93
intros
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    94
t0[intro]: "\<lbrakk>valid \<Gamma>; (x,\<tau>)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> I\<turnstile> IVar x : \<tau>"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    95
t1[intro]: "\<lbrakk>\<Gamma> I\<turnstile> e1 : \<tau>1\<rightarrow>\<tau>2; \<Gamma> I\<turnstile> e2 : \<tau>1\<rbrakk>\<Longrightarrow> \<Gamma> I\<turnstile> IApp e1 e2 : \<tau>2"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    96
t2[intro]: "\<lbrakk>x\<sharp>\<Gamma>;((x,\<tau>1)#\<Gamma>) I\<turnstile> t : \<tau>2\<rbrakk> \<Longrightarrow> \<Gamma> I\<turnstile> ILam [x].t : \<tau>1\<rightarrow>\<tau>2"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    97
t3[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> I\<turnstile> IUnit : DataI(OneI)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    98
t4[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> I\<turnstile> INat(n) : DataI(NatI)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
    99
t5[intro]: "\<Gamma> I\<turnstile> e : DataI(NatI) \<Longrightarrow> \<Gamma> I\<turnstile> ISucc(e) : DataI(NatI)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   100
t6[intro]: "\<lbrakk>\<Gamma> I\<turnstile> e : DataI(NatI)\<rbrakk> \<Longrightarrow> \<Gamma> I\<turnstile> IRef e : DataI (NatI)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   101
t7[intro]: "\<lbrakk>\<Gamma> I\<turnstile> e1 : DataI(NatI); \<Gamma> I\<turnstile> e2 : DataI(NatI)\<rbrakk> \<Longrightarrow> \<Gamma> I\<turnstile> e1\<mapsto>e2 : DataI(OneI)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   102
t8[intro]: "\<lbrakk>\<Gamma> I\<turnstile> e1 : DataI(NatI); \<Gamma> I\<turnstile> e2 : \<tau>\<rbrakk> \<Longrightarrow> \<Gamma> I\<turnstile> e1;;e2 : \<tau>"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   103
t9[intro]: "\<lbrakk>\<Gamma> I\<turnstile> e: DataI(NatI); \<Gamma> I\<turnstile> e1 : \<tau>; \<Gamma> I\<turnstile> e2 : \<tau>\<rbrakk> \<Longrightarrow> \<Gamma> I\<turnstile> Iif e e1 e2 : \<tau>"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   104
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   105
text {* capture-avoiding substitution *}
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   106
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   107
lemma perm_if:
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   108
  fixes pi::"'a prm"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   109
  shows "pi\<bullet>(if b then c1 else c2) = (if (pi\<bullet>b) then (pi\<bullet>c1) else (pi\<bullet>c2))"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   110
apply(simp add: perm_fun_def)
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   111
done
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   112
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   113
lemma eq_eqvt:
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   114
  fixes pi::"name prm"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   115
  and   x::"'a::pt_name"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   116
  shows "pi\<bullet>(x=y) = ((pi\<bullet>x)=(pi\<bullet>y))"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   117
  apply(simp add: perm_bool perm_bij)
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   118
  done
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   119
21543
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   120
consts
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   121
  subst :: "'a \<Rightarrow> name \<Rightarrow> 'a \<Rightarrow> 'a"  ("_[_::=_]" [100,100,100] 100)
20606
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   122
21543
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   123
nominal_primrec
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   124
  "(Var x)[y::=t'] = (if x=y then t' else (Var x))"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   125
  "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   126
  "\<lbrakk>x\<sharp>y; x\<sharp>t'\<rbrakk> \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   127
  "(Const n)[y::=t'] = Const n"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   128
  "(Pr e1 e2)[y::=t'] = Pr (e1[y::=t']) (e2[y::=t'])"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   129
  "(Fst e)[y::=t'] = Fst (e[y::=t'])"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   130
  "(Snd e)[y::=t'] = Snd (e[y::=t'])"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   131
  "(InL e)[y::=t'] = InL (e[y::=t'])"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   132
  "(InR e)[y::=t'] = InR (e[y::=t'])"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   133
  "\<lbrakk>z\<noteq>x; x\<sharp>y; x\<sharp>e; x\<sharp>e2; z\<sharp>y; z\<sharp>e; z\<sharp>e1; x\<sharp>t'; z\<sharp>t'\<rbrakk> \<Longrightarrow>
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   134
     (Case e of inl x \<rightarrow> e1 | inr z \<rightarrow> e2)[y::=t'] =
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   135
       (Case (e[y::=t']) of inl x \<rightarrow> (e1[y::=t']) | inr z \<rightarrow> (e2[y::=t']))"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   136
  apply (finite_guess add: fs_name1 perm_if eq_eqvt)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   137
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   138
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   139
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   140
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   141
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   142
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   143
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   144
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   145
  apply (finite_guess only: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   146
  apply perm_simp
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   147
  apply (simp add: supp_unit)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   148
  apply (rule TrueI)+
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   149
  apply (simp add: abs_fresh)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   150
  apply (simp_all add: abs_fresh)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   151
  apply (fresh_guess add: fs_name1 perm_if eq_eqvt)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   152
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   153
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   154
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   155
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   156
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   157
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   158
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   159
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   160
  apply (fresh_guess only: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   161
  apply perm_simp
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   162
  apply (fresh_guess add: fs_name1 perm_if eq_eqvt)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   163
  apply (fresh_guess add: fs_name1 perm_if eq_eqvt)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   164
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   165
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   166
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   167
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   168
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   169
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   170
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   171
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   172
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   173
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   174
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   175
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   176
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   177
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   178
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   179
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   180
  apply (fresh_guess only: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   181
  apply perm_simp
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   182
  apply (fresh_guess only: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   183
  apply perm_simp
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   184
  done
20606
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   185
21543
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   186
nominal_primrec (Isubst)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   187
  "(IVar x)[y::=t'] = (if x=y then t' else (IVar x))"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   188
  "(IApp t1 t2)[y::=t'] = IApp (t1[y::=t']) (t2[y::=t'])"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   189
  "\<lbrakk>x\<sharp>y; x\<sharp>t'\<rbrakk> \<Longrightarrow> (ILam [x].t)[y::=t'] = ILam [x].(t[y::=t'])"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   190
  "(INat n)[y::=t'] = INat n"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   191
  "(IUnit)[y::=t'] = IUnit"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   192
  "(ISucc e)[y::=t'] = ISucc (e[y::=t'])"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   193
  "(IAss e1 e2)[y::=t'] = IAss (e1[y::=t']) (e2[y::=t'])"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   194
  "(IRef e)[y::=t'] = IRef (e[y::=t'])"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   195
  "(ISeq e1 e2)[y::=t'] = ISeq (e1[y::=t']) (e2[y::=t'])"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   196
  "(Iif e e1 e2)[y::=t'] = Iif (e[y::=t']) (e1[y::=t']) (e2[y::=t'])"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   197
  apply (finite_guess add: fs_name1 perm_if eq_eqvt)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   198
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   199
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   200
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   201
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   202
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   203
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   204
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   205
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   206
  apply (finite_guess only: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   207
  apply perm_simp
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   208
  apply (simp add: supp_unit)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   209
  apply (rule TrueI)+
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   210
  apply (simp add: abs_fresh)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   211
  apply (fresh_guess add: fs_name1 perm_if eq_eqvt)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   212
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   213
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   214
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   215
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   216
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   217
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   218
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   219
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   220
  apply (fresh_guess only: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   221
  apply perm_simp
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   222
  done
20606
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   223
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   224
lemma Isubst_eqvt:
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   225
  fixes pi::"name prm"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   226
  and   t1::"trmI"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   227
  and   t2::"trmI"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   228
  and   x::"name"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   229
  shows "pi\<bullet>(t1[x::=t2]) = ((pi\<bullet>t1)[(pi\<bullet>x)::=(pi\<bullet>t2)])"
21543
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   230
  apply (nominal_induct t1 avoiding: x t2 rule: trmI.induct)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   231
  apply (simp_all add: Isubst.simps perm_if eq_eqvt fresh_bij)
20606
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   232
  done
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   233
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   234
lemma Isubst_supp:
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   235
  fixes t1::"trmI"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   236
  and   t2::"trmI"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   237
  and   x::"name"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   238
  shows "((supp (t1[x::=t2]))::name set) \<subseteq> (supp t2)\<union>((supp t1)-{x})"
21543
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   239
  apply (nominal_induct t1 avoiding: x t2 rule: trmI.induct)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   240
  apply (auto simp add: Isubst.simps trmI.supp supp_atm abs_supp supp_nat)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   241
  apply blast+
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   242
  done
20606
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   243
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   244
lemma Isubst_fresh:
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   245
  fixes x::"name"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   246
  and   y::"name"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   247
  and   t1::"trmI"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   248
  and   t2::"trmI"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   249
  assumes a: "x\<sharp>[y].t1" "x\<sharp>t2"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   250
  shows "x\<sharp>(t1[y::=t2])"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   251
using a
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   252
apply(auto simp add: fresh_def Isubst_supp)
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   253
apply(drule rev_subsetD)
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   254
apply(rule Isubst_supp)
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   255
apply(simp add: abs_supp)
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   256
done
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   257
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   258
text {* big-step evaluation for trms *}
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   259
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   260
consts
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   261
  big :: "(trm\<times>trm) set" 
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   262
syntax
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   263
  "_big_judge" :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<Down> _" [80,80] 80) 
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   264
translations
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   265
  "e1 \<Down> e2" \<rightleftharpoons> "(e1,e2) \<in> big"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   266
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   267
inductive big
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   268
intros
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   269
b0[intro]: "Lam [x].e \<Down> Lam [x].e"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   270
b1[intro]: "\<lbrakk>e1\<Down>Lam [x].e; e2\<Down>e2'; e[x::=e2']\<Down>e'\<rbrakk> \<Longrightarrow> App e1 e2 \<Down> e'"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   271
b2[intro]: "Const n \<Down> Const n"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   272
b3[intro]: "\<lbrakk>e1\<Down>e1'; e2\<Down>e2'\<rbrakk> \<Longrightarrow> Pr e1 e2 \<Down> Pr e1' e2'"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   273
b4[intro]: "e\<Down>Pr e1 e2 \<Longrightarrow> Fst e\<Down>e1"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   274
b5[intro]: "e\<Down>Pr e1 e2 \<Longrightarrow> Snd e\<Down>e2"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   275
b6[intro]: "e\<Down>e' \<Longrightarrow> InL e \<Down> InL e'"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   276
b7[intro]: "e\<Down>e' \<Longrightarrow> InR e \<Down> InR e'"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   277
b8[intro]: "\<lbrakk>e\<Down>InL e'; e1[x::=e']\<Down>e''\<rbrakk> \<Longrightarrow> Case e of inl x1 \<rightarrow> e1 | inr x2 \<rightarrow> e2 \<Down> e''"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   278
b9[intro]: "\<lbrakk>e\<Down>InR e'; e2[x::=e']\<Down>e''\<rbrakk> \<Longrightarrow> Case e of inl x1 \<rightarrow> e1 | inr x2 \<rightarrow> e2 \<Down> e''"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   279
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   280
consts
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   281
  Ibig :: "(((nat\<Rightarrow>nat)\<times>trmI)\<times>((nat\<Rightarrow>nat)\<times>trmI)) set" 
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   282
syntax
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   283
  "_Ibig_judge" :: "((nat\<Rightarrow>nat)\<times>trmI)\<Rightarrow>((nat\<Rightarrow>nat)\<times>trmI)\<Rightarrow>bool" ("_ I\<Down> _" [80,80] 80) 
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   284
translations
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   285
  "(m,e1) I\<Down> (m',e2)" \<rightleftharpoons> "((m,e1),(m',e2)) \<in> Ibig"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   286
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   287
inductive Ibig
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   288
intros
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   289
m0[intro]: "(m,ILam [x].e) I\<Down> (m,ILam [x].e)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   290
m1[intro]: "\<lbrakk>(m,e1)I\<Down>(m',ILam [x].e); (m',e2)I\<Down>(m'',e3); (m'',e[x::=e3])I\<Down>(m''',e4)\<rbrakk> 
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   291
            \<Longrightarrow> (m,IApp e1 e2) I\<Down> (m''',e4)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   292
m2[intro]: "(m,IUnit) I\<Down> (m,IUnit)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   293
m3[intro]: "(m,INat(n))I\<Down>(m,INat(n))"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   294
m4[intro]: "(m,e)I\<Down>(m',INat(n)) \<Longrightarrow> (m,ISucc(e))I\<Down>(m',INat(n+1))"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   295
m5[intro]: "(m,e)I\<Down>(m',INat(n)) \<Longrightarrow> (m,IRef(e))I\<Down>(m',INat(m' n))"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   296
m6[intro]: "\<lbrakk>(m,e1)I\<Down>(m',INat(n1)); (m',e2)I\<Down>(m'',INat(n2))\<rbrakk> \<Longrightarrow> (m,e1\<mapsto>e2)I\<Down>(m''(n1:=n2),IUnit)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   297
m7[intro]: "\<lbrakk>(m,e1)I\<Down>(m',IUnit); (m',e2)I\<Down>(m'',e)\<rbrakk> \<Longrightarrow> (m,e1;;e2)I\<Down>(m'',e)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   298
m8[intro]: "\<lbrakk>(m,e)I\<Down>(m',INat(n)); n\<noteq>0; (m',e1)I\<Down>(m'',e)\<rbrakk> \<Longrightarrow> (m,Iif e e1 e2)I\<Down>(m'',e)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   299
m9[intro]: "\<lbrakk>(m,e)I\<Down>(m',INat(0)); (m',e2)I\<Down>(m'',e)\<rbrakk> \<Longrightarrow> (m,Iif e e1 e2)I\<Down>(m'',e)"
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   300
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   301
text {* Translation functions *}
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   302
21543
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   303
consts trans :: "trm \<Rightarrow> trmI" 
20606
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   304
21543
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   305
nominal_primrec
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   306
  "trans (Var x) = (IVar x)"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   307
  "trans (App e1 e2) = IApp (trans e1) (trans e2)"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   308
  "trans (Lam [x].e) = ILam [x].(trans e)"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   309
  "trans (Const n) = INat n"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   310
  "trans (Pr e1 e2) = 
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   311
       (let limit = IRef(INat 0) in 
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   312
        let v1 = (trans e1) in 
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   313
        let v2 = (trans e2) in 
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   314
        (((ISucc limit)\<mapsto>v1);;(ISucc(ISucc limit)\<mapsto>v2));;(INat 0 \<mapsto> ISucc(ISucc(limit))))"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   315
  "trans (Fst e) = IRef (ISucc (trans e))"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   316
  "trans (Snd e) = IRef (ISucc (ISucc (trans e)))"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   317
  "trans (InL e) = 
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   318
        (let limit = IRef(INat 0) in 
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   319
         let v = (trans e) in 
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   320
         (((ISucc limit)\<mapsto>INat(0));;(ISucc(ISucc limit)\<mapsto>v));;(INat 0 \<mapsto> ISucc(ISucc(limit))))"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   321
  "trans (InR e) = 
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   322
        (let limit = IRef(INat 0) in 
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   323
         let v = (trans e) in 
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   324
         (((ISucc limit)\<mapsto>INat(1));;(ISucc(ISucc limit)\<mapsto>v));;(INat 0 \<mapsto> ISucc(ISucc(limit))))"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   325
  "\<lbrakk>x2\<noteq>x1; x1\<sharp>e; x1\<sharp>e2; x2\<sharp>e; x2\<sharp>e1\<rbrakk> \<Longrightarrow> 
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   326
   trans (Case e of inl x1 \<rightarrow> e1 | inr x2 \<rightarrow> e2) =
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   327
       (let v = (trans e) in
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   328
        let v1 = (trans e1) in
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   329
        let v2 = (trans e2) in 
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   330
        Iif (IRef (ISucc v)) (v2[x2::=IRef (ISucc (ISucc v))]) (v1[x1::=IRef (ISucc (ISucc v))]))"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   331
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   332
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   333
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   334
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   335
  apply (finite_guess add: fs_name1 Let_def perm_nat_def)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   336
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   337
  apply (finite_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   338
  apply (finite_guess add: fs_name1 Let_def perm_nat_def)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   339
  apply (finite_guess add: fs_name1 Let_def perm_nat_def)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   340
  apply (finite_guess add: fs_name1 Let_def Isubst_eqvt)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   341
  apply (simp add: supp_unit)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   342
  apply (rule TrueI)+
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   343
  apply (simp add: abs_fresh)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   344
  apply (simp_all add: abs_fresh Isubst_fresh)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   345
  apply (fresh_guess add: fs_name1 perm_if eq_eqvt)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   346
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   347
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   348
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   349
  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   350
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   351
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   352
  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   353
  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   354
  apply (fresh_guess add: fs_name1 Let_def Isubst_eqvt)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   355
  apply (fresh_guess add: fs_name1 perm_if eq_eqvt)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   356
  apply (fresh_guess add: fs_name1 perm_if eq_eqvt)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   357
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   358
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   359
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   360
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   361
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   362
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   363
  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   364
  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   365
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   366
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   367
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   368
  apply (fresh_guess add: fs_name1)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   369
  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   370
  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   371
  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   372
  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   373
  apply (fresh_guess add: fs_name1 Let_def Isubst_eqvt)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   374
  apply (fresh_guess add: fs_name1 Let_def Isubst_eqvt)
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   375
  done
20606
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   376
21543
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   377
consts trans_type :: "ty \<Rightarrow> tyI"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   378
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   379
nominal_primrec
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   380
  "trans_type (Data \<sigma>) = DataI(NatI)"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   381
  "trans_type (\<tau>1\<rightarrow>\<tau>2) = (trans_type \<tau>1)\<rightarrow>(trans_type \<tau>2)"
e855f25df0c8 Adapted to new nominal_primrec command.
berghofe
parents: 21087
diff changeset
   382
  by (rule TrueI)+
20606
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   383
fd9b0b78a7d3 this file contains a compile-challenge suggested by Adam Chlipala;
urbanc
parents:
diff changeset
   384
end