# HG changeset patch # User wenzelm # Date 1275919235 -7200 # Node ID 74fb4f03bb51886f597156d350e1e21d70b32d9e # Parent 1255ba99ed1e6381c0546e515963783753440230 recovered some untested theories; diff -r 1255ba99ed1e -r 74fb4f03bb51 src/HOL/Nominal/Examples/CK_Machine.thy --- a/src/HOL/Nominal/Examples/CK_Machine.thy Mon Jun 07 13:20:05 2010 +0200 +++ b/src/HOL/Nominal/Examples/CK_Machine.thy Mon Jun 07 16:00:35 2010 +0200 @@ -492,7 +492,7 @@ shows "(\@\) \ e[x::=e'] : T" using a b proof (nominal_induct \'\"\@[(x,T')]@\" e T avoiding: x e' \ rule: typing.strong_induct) - case (t_VAR \' y T x e' \) + case (t_VAR y T x e' \) then have a1: "valid (\@[(x,T')]@\)" and a2: "(y,T) \ set (\@[(x,T')]@\)" and a3: "\ \ e' : T'" by simp_all diff -r 1255ba99ed1e -r 74fb4f03bb51 src/HOL/Nominal/Examples/Nominal_Examples.thy --- a/src/HOL/Nominal/Examples/Nominal_Examples.thy Mon Jun 07 13:20:05 2010 +0200 +++ b/src/HOL/Nominal/Examples/Nominal_Examples.thy Mon Jun 07 16:00:35 2010 +0200 @@ -4,23 +4,25 @@ theory Nominal_Examples imports + CK_Machine CR CR_Takahashi Class3 Compile + Contexts + Crary Fsub Height Lambda_mu + LocalWeakening + Pattern SN - Weakening - Crary SOS - LocalWeakening + Standardization Support - Contexts - Standardization + Type_Preservation W - Pattern + Weakening begin end diff -r 1255ba99ed1e -r 74fb4f03bb51 src/HOL/Nominal/Examples/Type_Preservation.thy --- a/src/HOL/Nominal/Examples/Type_Preservation.thy Mon Jun 07 13:20:05 2010 +0200 +++ b/src/HOL/Nominal/Examples/Type_Preservation.thy Mon Jun 07 16:00:35 2010 +0200 @@ -142,7 +142,7 @@ shows "(\@\) \ e[x::=e'] : T" using a b proof (nominal_induct \'\"\@[(x,T')]@\" e T avoiding: x e' \ rule: typing.strong_induct) - case (t_Var \' y T x e' \) + case (t_Var y T x e' \) then have a1: "valid (\@[(x,T')]@\)" and a2: "(y,T) \ set (\@[(x,T')]@\)" and a3: "\ \ e' : T'" by simp_all