| author | wenzelm | 
| Mon, 27 Nov 2017 10:21:52 +0100 | |
| changeset 67093 | 835a2ab92c3d | 
| parent 66453 | cc19f7ca2ed6 | 
| child 80148 | b156869b826a | 
| permissions | -rw-r--r-- | 
| 22829 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 1 | (* The definitions for a challenge suggested by Adam Chlipala *) | 
| 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 | theory Compile | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
63167diff
changeset | 4 | imports "HOL-Nominal.Nominal" | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 5 | begin | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 6 | |
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 7 | atom_decl name | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 8 | |
| 22829 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 9 | nominal_datatype data = | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 10 | DNat | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 11 | | DProd "data" "data" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 12 | | DSum "data" "data" | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 13 | |
| 22829 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 14 | nominal_datatype ty = | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 15 | Data "data" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 16 |   | Arrow "ty" "ty" ("_\<rightarrow>_" [100,100] 100)
 | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 17 | |
| 22829 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 18 | nominal_datatype trm = | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 19 | Var "name" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 20 |   | Lam "\<guillemotleft>name\<guillemotright>trm" ("Lam [_]._" [100,100] 100)
 | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 21 | | App "trm" "trm" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 22 | | Const "nat" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 23 | | Pr "trm" "trm" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 24 | | Fst "trm" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 25 | | Snd "trm" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 26 | | InL "trm" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 27 | | InR "trm" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 28 | | Case "trm" "\<guillemotleft>name\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 29 |           ("Case _ of inl _ \<rightarrow> _ | inr _ \<rightarrow> _" [100,100,100,100,100] 100)
 | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 30 | |
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 31 | nominal_datatype dataI = OneI | NatI | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 32 | |
| 22829 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 33 | nominal_datatype tyI = | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 34 | DataI "dataI" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 35 |   | ArrowI "tyI" "tyI" ("_\<rightarrow>_" [100,100] 100)
 | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 36 | |
| 22829 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 37 | nominal_datatype trmI = | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 38 | IVar "name" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 39 |   | ILam "\<guillemotleft>name\<guillemotright>trmI" ("ILam [_]._" [100,100] 100)
 | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 40 | | IApp "trmI" "trmI" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 41 | | IUnit | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 42 | | INat "nat" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 43 | | ISucc "trmI" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 44 |   | IAss "trmI" "trmI" ("_\<mapsto>_" [100,100] 100)
 | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 45 | | IRef "trmI" | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 46 |   | ISeq "trmI" "trmI" ("_;;_" [100,100] 100)
 | 
| 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 47 | | Iif "trmI" "trmI" "trmI" | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 48 | |
| 63167 | 49 | text \<open>valid contexts\<close> | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 50 | |
| 23760 | 51 | inductive | 
| 22829 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 52 | valid :: "(name\<times>'a::pt_name) list \<Rightarrow> bool" | 
| 22271 | 53 | where | 
| 54 | v1[intro]: "valid []" | |
| 55 | | 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 | 56 | |
| 63167 | 57 | text \<open>typing judgements for trms\<close> | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 58 | |
| 23760 | 59 | inductive | 
| 22829 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 60 |   typing :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [80,80,80] 80)
 | 
| 22271 | 61 | where | 
| 62 | t0[intro]: "\<lbrakk>valid \<Gamma>; (x,\<tau>)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var x : \<tau>" | |
| 63 | | 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" | |
| 64 | | 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" | |
| 65 | | t3[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Const n : Data(DNat)" | |
| 66 | | 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)" | |
| 67 | | t5[intro]: "\<lbrakk>\<Gamma> \<turnstile> e : Data(DProd \<sigma>1 \<sigma>2)\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Fst e : Data(\<sigma>1)" | |
| 68 | | t6[intro]: "\<lbrakk>\<Gamma> \<turnstile> e : Data(DProd \<sigma>1 \<sigma>2)\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Snd e : Data(\<sigma>2)" | |
| 69 | | t7[intro]: "\<lbrakk>\<Gamma> \<turnstile> e : Data(\<sigma>1)\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> InL e : Data(DSum \<sigma>1 \<sigma>2)" | |
| 70 | | t8[intro]: "\<lbrakk>\<Gamma> \<turnstile> e : Data(\<sigma>2)\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> InR e : Data(DSum \<sigma>1 \<sigma>2)" | |
| 71 | | t9[intro]: "\<lbrakk>x1\<sharp>\<Gamma>; x2\<sharp>\<Gamma>; \<Gamma> \<turnstile> e: Data(DSum \<sigma>1 \<sigma>2); | |
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 72 | ((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 | 73 | \<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 | 74 | |
| 63167 | 75 | text \<open>typing judgements for Itrms\<close> | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 76 | |
| 23760 | 77 | inductive | 
| 22829 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 78 |   Ityping :: "(name\<times>tyI) list\<Rightarrow>trmI\<Rightarrow>tyI\<Rightarrow>bool" (" _ I\<turnstile> _ : _ " [80,80,80] 80)
 | 
| 22271 | 79 | where | 
| 80 | t0[intro]: "\<lbrakk>valid \<Gamma>; (x,\<tau>)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> I\<turnstile> IVar x : \<tau>" | |
| 81 | | 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" | |
| 82 | | 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" | |
| 83 | | t3[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> I\<turnstile> IUnit : DataI(OneI)" | |
| 84 | | t4[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> I\<turnstile> INat(n) : DataI(NatI)" | |
| 85 | | t5[intro]: "\<Gamma> I\<turnstile> e : DataI(NatI) \<Longrightarrow> \<Gamma> I\<turnstile> ISucc(e) : DataI(NatI)" | |
| 86 | | t6[intro]: "\<lbrakk>\<Gamma> I\<turnstile> e : DataI(NatI)\<rbrakk> \<Longrightarrow> \<Gamma> I\<turnstile> IRef e : DataI (NatI)" | |
| 87 | | 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)" | |
| 88 | | t8[intro]: "\<lbrakk>\<Gamma> I\<turnstile> e1 : DataI(NatI); \<Gamma> I\<turnstile> e2 : \<tau>\<rbrakk> \<Longrightarrow> \<Gamma> I\<turnstile> e1;;e2 : \<tau>" | |
| 89 | | 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>" | |
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 90 | |
| 63167 | 91 | text \<open>capture-avoiding substitution\<close> | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 92 | |
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 93 | class subst = | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 94 |   fixes 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 | 95 | |
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 96 | instantiation trm :: subst | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 97 | begin | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 98 | |
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 99 | nominal_primrec subst_trm | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 100 | where | 
| 21543 | 101 | "(Var x)[y::=t'] = (if x=y then t' else (Var x))" | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 102 | | "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 103 | | "\<lbrakk>x\<sharp>y; x\<sharp>t'\<rbrakk> \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 104 | | "(Const n)[y::=t'] = Const n" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 105 | | "(Pr e1 e2)[y::=t'] = Pr (e1[y::=t']) (e2[y::=t'])" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 106 | | "(Fst e)[y::=t'] = Fst (e[y::=t'])" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 107 | | "(Snd e)[y::=t'] = Snd (e[y::=t'])" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 108 | | "(InL e)[y::=t'] = InL (e[y::=t'])" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 109 | | "(InR e)[y::=t'] = InR (e[y::=t'])" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 110 | | "\<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> | 
| 21543 | 111 | (Case e of inl x \<rightarrow> e1 | inr z \<rightarrow> e2)[y::=t'] = | 
| 112 | (Case (e[y::=t']) of inl x \<rightarrow> (e1[y::=t']) | inr z \<rightarrow> (e2[y::=t']))" | |
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 113 | apply(finite_guess)+ | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 114 | apply(rule TrueI)+ | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 115 | apply(simp add: abs_fresh)+ | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 116 | apply(fresh_guess)+ | 
| 21543 | 117 | done | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 118 | |
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 119 | instance .. | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 120 | |
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 121 | end | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 122 | |
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 123 | instantiation trmI :: subst | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 124 | begin | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 125 | |
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 126 | nominal_primrec subst_trmI | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 127 | where | 
| 21543 | 128 | "(IVar x)[y::=t'] = (if x=y then t' else (IVar x))" | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 129 | | "(IApp t1 t2)[y::=t'] = IApp (t1[y::=t']) (t2[y::=t'])" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 130 | | "\<lbrakk>x\<sharp>y; x\<sharp>t'\<rbrakk> \<Longrightarrow> (ILam [x].t)[y::=t'] = ILam [x].(t[y::=t'])" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 131 | | "(INat n)[y::=t'] = INat n" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 132 | | "(IUnit)[y::=t'] = IUnit" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 133 | | "(ISucc e)[y::=t'] = ISucc (e[y::=t'])" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 134 | | "(IAss e1 e2)[y::=t'] = IAss (e1[y::=t']) (e2[y::=t'])" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 135 | | "(IRef e)[y::=t'] = IRef (e[y::=t'])" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 136 | | "(ISeq e1 e2)[y::=t'] = ISeq (e1[y::=t']) (e2[y::=t'])" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 137 | | "(Iif e e1 e2)[y::=t'] = Iif (e[y::=t']) (e1[y::=t']) (e2[y::=t'])" | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 138 | apply(finite_guess)+ | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 139 | apply(rule TrueI)+ | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 140 | apply(simp add: abs_fresh)+ | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 141 | apply(fresh_guess)+ | 
| 21543 | 142 | done | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 143 | |
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 144 | instance .. | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 145 | |
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 146 | end | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 147 | |
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 148 | lemma Isubst_eqvt[eqvt]: | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 149 | fixes pi::"name prm" | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 150 | and t1::"trmI" | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 151 | and t2::"trmI" | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 152 | and x::"name" | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 153 | shows "pi\<bullet>(t1[x::=t2]) = ((pi\<bullet>t1)[(pi\<bullet>x)::=(pi\<bullet>t2)])" | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
23760diff
changeset | 154 | apply (nominal_induct t1 avoiding: x t2 rule: trmI.strong_induct) | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 155 | apply (simp_all add: subst_trmI.simps eqvts fresh_bij) | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 156 | done | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 157 | |
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 158 | lemma Isubst_supp: | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 159 | fixes t1::"trmI" | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 160 | and t2::"trmI" | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 161 | and x::"name" | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 162 |   shows "((supp (t1[x::=t2]))::name set) \<subseteq> (supp t2)\<union>((supp t1)-{x})"
 | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
23760diff
changeset | 163 | apply (nominal_induct t1 avoiding: x t2 rule: trmI.strong_induct) | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 164 | apply (auto simp add: subst_trmI.simps trmI.supp supp_atm abs_supp supp_nat) | 
| 21543 | 165 | apply blast+ | 
| 166 | done | |
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 167 | |
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 168 | lemma Isubst_fresh: | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 169 | fixes x::"name" | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 170 | and y::"name" | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 171 | and t1::"trmI" | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 172 | and t2::"trmI" | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 173 | assumes a: "x\<sharp>[y].t1" "x\<sharp>t2" | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 174 | shows "x\<sharp>(t1[y::=t2])" | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 175 | using a | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 176 | apply(auto simp add: fresh_def Isubst_supp) | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 177 | apply(drule rev_subsetD) | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 178 | apply(rule Isubst_supp) | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 179 | apply(simp add: abs_supp) | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 180 | done | 
| 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 181 | |
| 63167 | 182 | text \<open>big-step evaluation for trms\<close> | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 183 | |
| 23760 | 184 | inductive | 
| 22829 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 185 |   big :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<Down> _" [80,80] 80)
 | 
| 22271 | 186 | where | 
| 187 | b0[intro]: "Lam [x].e \<Down> Lam [x].e" | |
| 188 | | b1[intro]: "\<lbrakk>e1\<Down>Lam [x].e; e2\<Down>e2'; e[x::=e2']\<Down>e'\<rbrakk> \<Longrightarrow> App e1 e2 \<Down> e'" | |
| 189 | | b2[intro]: "Const n \<Down> Const n" | |
| 190 | | b3[intro]: "\<lbrakk>e1\<Down>e1'; e2\<Down>e2'\<rbrakk> \<Longrightarrow> Pr e1 e2 \<Down> Pr e1' e2'" | |
| 191 | | b4[intro]: "e\<Down>Pr e1 e2 \<Longrightarrow> Fst e\<Down>e1" | |
| 192 | | b5[intro]: "e\<Down>Pr e1 e2 \<Longrightarrow> Snd e\<Down>e2" | |
| 193 | | b6[intro]: "e\<Down>e' \<Longrightarrow> InL e \<Down> InL e'" | |
| 194 | | b7[intro]: "e\<Down>e' \<Longrightarrow> InR e \<Down> InR e'" | |
| 195 | | 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''" | |
| 196 | | 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''" | |
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 197 | |
| 23760 | 198 | inductive | 
| 22829 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 urbanc parents: 
22541diff
changeset | 199 |   Ibig :: "((nat\<Rightarrow>nat)\<times>trmI)\<Rightarrow>((nat\<Rightarrow>nat)\<times>trmI)\<Rightarrow>bool" ("_ I\<Down> _" [80,80] 80)
 | 
| 22271 | 200 | where | 
| 201 | m0[intro]: "(m,ILam [x].e) I\<Down> (m,ILam [x].e)" | |
| 202 | | 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> | |
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 203 | \<Longrightarrow> (m,IApp e1 e2) I\<Down> (m''',e4)" | 
| 22271 | 204 | | m2[intro]: "(m,IUnit) I\<Down> (m,IUnit)" | 
| 205 | | m3[intro]: "(m,INat(n))I\<Down>(m,INat(n))" | |
| 206 | | m4[intro]: "(m,e)I\<Down>(m',INat(n)) \<Longrightarrow> (m,ISucc(e))I\<Down>(m',INat(n+1))" | |
| 207 | | m5[intro]: "(m,e)I\<Down>(m',INat(n)) \<Longrightarrow> (m,IRef(e))I\<Down>(m',INat(m' n))" | |
| 208 | | 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)" | |
| 209 | | m7[intro]: "\<lbrakk>(m,e1)I\<Down>(m',IUnit); (m',e2)I\<Down>(m'',e)\<rbrakk> \<Longrightarrow> (m,e1;;e2)I\<Down>(m'',e)" | |
| 210 | | 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)" | |
| 211 | | 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)" | |
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 212 | |
| 63167 | 213 | text \<open>Translation functions\<close> | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 214 | |
| 21543 | 215 | nominal_primrec | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 216 | trans :: "trm \<Rightarrow> trmI" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 217 | where | 
| 21543 | 218 | "trans (Var x) = (IVar x)" | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 219 | | "trans (App e1 e2) = IApp (trans e1) (trans e2)" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 220 | | "trans (Lam [x].e) = ILam [x].(trans e)" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 221 | | "trans (Const n) = INat n" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 222 | | "trans (Pr e1 e2) = | 
| 21543 | 223 | (let limit = IRef(INat 0) in | 
| 224 | let v1 = (trans e1) in | |
| 225 | let v2 = (trans e2) in | |
| 226 | (((ISucc limit)\<mapsto>v1);;(ISucc(ISucc limit)\<mapsto>v2));;(INat 0 \<mapsto> ISucc(ISucc(limit))))" | |
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 227 | | "trans (Fst e) = IRef (ISucc (trans e))" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 228 | | "trans (Snd e) = IRef (ISucc (ISucc (trans e)))" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 229 | | "trans (InL e) = | 
| 21543 | 230 | (let limit = IRef(INat 0) in | 
| 231 | let v = (trans e) in | |
| 232 | (((ISucc limit)\<mapsto>INat(0));;(ISucc(ISucc limit)\<mapsto>v));;(INat 0 \<mapsto> ISucc(ISucc(limit))))" | |
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 233 | | "trans (InR e) = | 
| 21543 | 234 | (let limit = IRef(INat 0) in | 
| 235 | let v = (trans e) in | |
| 236 | (((ISucc limit)\<mapsto>INat(1));;(ISucc(ISucc limit)\<mapsto>v));;(INat 0 \<mapsto> ISucc(ISucc(limit))))" | |
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 237 | | "\<lbrakk>x2\<noteq>x1; x1\<sharp>e; x1\<sharp>e2; x2\<sharp>e; x2\<sharp>e1\<rbrakk> \<Longrightarrow> | 
| 21543 | 238 | trans (Case e of inl x1 \<rightarrow> e1 | inr x2 \<rightarrow> e2) = | 
| 239 | (let v = (trans e) in | |
| 240 | let v1 = (trans e1) in | |
| 241 | let v2 = (trans e2) in | |
| 242 | Iif (IRef (ISucc v)) (v2[x2::=IRef (ISucc (ISucc v))]) (v1[x1::=IRef (ISucc (ISucc v))]))" | |
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 243 | apply(finite_guess add: Let_def)+ | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 244 | apply(rule TrueI)+ | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 245 | apply(simp add: abs_fresh Isubst_fresh)+ | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22271diff
changeset | 246 | apply(fresh_guess add: Let_def)+ | 
| 21543 | 247 | done | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 248 | |
| 21543 | 249 | nominal_primrec | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 250 | trans_type :: "ty \<Rightarrow> tyI" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 251 | where | 
| 21543 | 252 | "trans_type (Data \<sigma>) = DataI(NatI)" | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 253 | | "trans_type (\<tau>1\<rightarrow>\<tau>2) = (trans_type \<tau>1)\<rightarrow>(trans_type \<tau>2)" | 
| 21543 | 254 | by (rule TrueI)+ | 
| 20606 
fd9b0b78a7d3
this file contains a compile-challenge suggested by Adam Chlipala;
 urbanc parents: diff
changeset | 255 | |
| 62390 | 256 | end |