# HG changeset patch # User berghofe # Date 1229184789 -3600 # Node ID 66fe138979f4977e86df43f65e14caa56e8818a8 # Parent 1b8e021e8c1f324038666d4dd2ab4cea2563aafd# Parent 933145f5a9ba960c869ad9b77490dd996c165dfc merged diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/Examples/CK_Machine.thy --- a/src/HOL/Nominal/Examples/CK_Machine.thy Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/Examples/CK_Machine.thy Sat Dec 13 17:13:09 2008 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory CK_Machine imports "../Nominal" begin @@ -41,21 +39,21 @@ section {* Capture-Avoiding Substitution *} -consts subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [100,100,100] 100) - nominal_primrec + subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [100,100,100] 100) +where "(VAR x)[y::=s] = (if x=y then s else (VAR x))" - "(APP t\<^isub>1 t\<^isub>2)[y::=s] = APP (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])" - "x\(y,s) \ (LAM [x].t)[y::=s] = LAM [x].(t[y::=s])" - "(NUM n)[y::=s] = NUM n" - "(t\<^isub>1 -- t\<^isub>2)[y::=s] = (t\<^isub>1[y::=s]) -- (t\<^isub>2[y::=s])" - "(t\<^isub>1 ++ t\<^isub>2)[y::=s] = (t\<^isub>1[y::=s]) ++ (t\<^isub>2[y::=s])" - "x\(y,s) \ (FIX [x].t)[y::=s] = FIX [x].(t[y::=s])" - "TRUE[y::=s] = TRUE" - "FALSE[y::=s] = FALSE" - "(IF t1 t2 t3)[y::=s] = IF (t1[y::=s]) (t2[y::=s]) (t3[y::=s])" - "(ZET t)[y::=s] = ZET (t[y::=s])" - "(EQI t1 t2)[y::=s] = EQI (t1[y::=s]) (t2[y::=s])" +| "(APP t\<^isub>1 t\<^isub>2)[y::=s] = APP (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])" +| "x\(y,s) \ (LAM [x].t)[y::=s] = LAM [x].(t[y::=s])" +| "(NUM n)[y::=s] = NUM n" +| "(t\<^isub>1 -- t\<^isub>2)[y::=s] = (t\<^isub>1[y::=s]) -- (t\<^isub>2[y::=s])" +| "(t\<^isub>1 ++ t\<^isub>2)[y::=s] = (t\<^isub>1[y::=s]) ++ (t\<^isub>2[y::=s])" +| "x\(y,s) \ (FIX [x].t)[y::=s] = FIX [x].(t[y::=s])" +| "TRUE[y::=s] = TRUE" +| "FALSE[y::=s] = FALSE" +| "(IF t1 t2 t3)[y::=s] = IF (t1[y::=s]) (t2[y::=s]) (t3[y::=s])" +| "(ZET t)[y::=s] = ZET (t[y::=s])" +| "(EQI t1 t2)[y::=s] = EQI (t1[y::=s]) (t2[y::=s])" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh)+ diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/Examples/CR_Takahashi.thy --- a/src/HOL/Nominal/Examples/CR_Takahashi.thy Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/Examples/CR_Takahashi.thy Sat Dec 13 17:13:09 2008 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - (* Authors: Christian Urban and Mathilde Arnaud *) (* *) (* A formalisation of the Church-Rosser proof by Masako Takahashi.*) @@ -20,12 +18,12 @@ | App "lam" "lam" | Lam "\name\lam" ("Lam [_]._" [100,100] 100) -consts subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [100,100,100] 100) - nominal_primrec + subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [100,100,100] 100) +where "(Var x)[y::=s] = (if x=y then s else (Var x))" - "(App t\<^isub>1 t\<^isub>2)[y::=s] = App (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])" - "x\(y,s) \ (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])" +| "(App t\<^isub>1 t\<^isub>2)[y::=s] = App (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])" +| "x\(y,s) \ (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh) diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/Examples/Class.thy --- a/src/HOL/Nominal/Examples/Class.thy Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/Examples/Class.thy Sat Dec 13 17:13:09 2008 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory Class imports "../Nominal" begin @@ -17,16 +15,22 @@ | OR "ty" "ty" ("_ OR _" [100,100] 100) | IMP "ty" "ty" ("_ IMP _" [100,100] 100) -instance ty :: size .. - -nominal_primrec +instantiation ty :: size +begin + +nominal_primrec size_ty +where "size (PR s) = (1::nat)" - "size (NOT T) = 1 + size T" - "size (T1 AND T2) = 1 + size T1 + size T2" - "size (T1 OR T2) = 1 + size T1 + size T2" - "size (T1 IMP T2) = 1 + size T1 + size T2" +| "size (NOT T) = 1 + size T" +| "size (T1 AND T2) = 1 + size T1 + size T2" +| "size (T1 OR T2) = 1 + size T1 + size T2" +| "size (T1 IMP T2) = 1 + size T1 + size T2" by (rule TrueI)+ +instance .. + +end + lemma ty_cases: fixes T::ty shows "(\s. T=PR s) \ (\T'. T=NOT T') \ (\S U. T=S OR U) \ (\S U. T=S AND U) \ (\S U. T=S IMP U)" @@ -66,25 +70,23 @@ text {* renaming functions *} -consts - nrename :: "trm \ name \ name \ trm" ("_[_\n>_]" [100,100,100] 100) +nominal_primrec (freshness_context: "(d::coname,e::coname)") crename :: "trm \ coname \ coname \ trm" ("_[_\c>_]" [100,100,100] 100) - -nominal_primrec (freshness_context: "(d::coname,e::coname)") +where "(Ax x a)[d\c>e] = (if a=d then Ax x e else Ax x a)" - "\a\(d,e,N);x\M\ \ (Cut .M (x).N)[d\c>e] = Cut .(M[d\c>e]) (x).(N[d\c>e])" - "(NotR (x).M a)[d\c>e] = (if a=d then NotR (x).(M[d\c>e]) e else NotR (x).(M[d\c>e]) a)" - "a\(d,e) \ (NotL .M x)[d\c>e] = (NotL .(M[d\c>e]) x)" - "\a\(d,e,N,c);b\(d,e,M,c);b\a\ \ (AndR .M .N c)[d\c>e] = +| "\a\(d,e,N);x\M\ \ (Cut .M (x).N)[d\c>e] = Cut .(M[d\c>e]) (x).(N[d\c>e])" +| "(NotR (x).M a)[d\c>e] = (if a=d then NotR (x).(M[d\c>e]) e else NotR (x).(M[d\c>e]) a)" +| "a\(d,e) \ (NotL .M x)[d\c>e] = (NotL .(M[d\c>e]) x)" +| "\a\(d,e,N,c);b\(d,e,M,c);b\a\ \ (AndR .M .N c)[d\c>e] = (if c=d then AndR .(M[d\c>e]) .(N[d \c>e]) e else AndR .(M[d\c>e]) .(N[d\c>e]) c)" - "x\y \ (AndL1 (x).M y)[d\c>e] = AndL1 (x).(M[d\c>e]) y" - "x\y \ (AndL2 (x).M y)[d\c>e] = AndL2 (x).(M[d\c>e]) y" - "a\(d,e,b) \ (OrR1 .M b)[d\c>e] = (if b=d then OrR1 .(M[d\c>e]) e else OrR1 .(M[d\c>e]) b)" - "a\(d,e,b) \ (OrR2 .M b)[d\c>e] = (if b=d then OrR2 .(M[d\c>e]) e else OrR2 .(M[d\c>e]) b)" - "\x\(N,z);y\(M,z);y\x\ \ (OrL (x).M (y).N z)[d\c>e] = OrL (x).(M[d\c>e]) (y).(N[d\c>e]) z" - "a\(d,e,b) \ (ImpR (x)..M b)[d\c>e] = +| "x\y \ (AndL1 (x).M y)[d\c>e] = AndL1 (x).(M[d\c>e]) y" +| "x\y \ (AndL2 (x).M y)[d\c>e] = AndL2 (x).(M[d\c>e]) y" +| "a\(d,e,b) \ (OrR1 .M b)[d\c>e] = (if b=d then OrR1 .(M[d\c>e]) e else OrR1 .(M[d\c>e]) b)" +| "a\(d,e,b) \ (OrR2 .M b)[d\c>e] = (if b=d then OrR2 .(M[d\c>e]) e else OrR2 .(M[d\c>e]) b)" +| "\x\(N,z);y\(M,z);y\x\ \ (OrL (x).M (y).N z)[d\c>e] = OrL (x).(M[d\c>e]) (y).(N[d\c>e]) z" +| "a\(d,e,b) \ (ImpR (x)..M b)[d\c>e] = (if b=d then ImpR (x)..(M[d\c>e]) e else ImpR (x)..(M[d\c>e]) b)" - "\a\(d,e,N);x\(M,y)\ \ (ImpL .M (x).N y)[d\c>e] = ImpL .(M[d\c>e]) (x).(N[d\c>e]) y" +| "\a\(d,e,N);x\(M,y)\ \ (ImpL .M (x).N y)[d\c>e] = ImpL .(M[d\c>e]) (x).(N[d\c>e]) y" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh abs_supp fin_supp)+ @@ -92,19 +94,21 @@ done nominal_primrec (freshness_context: "(u::name,v::name)") + nrename :: "trm \ name \ name \ trm" ("_[_\n>_]" [100,100,100] 100) +where "(Ax x a)[u\n>v] = (if x=u then Ax v a else Ax x a)" - "\a\N;x\(u,v,M)\ \ (Cut .M (x).N)[u\n>v] = Cut .(M[u\n>v]) (x).(N[u\n>v])" - "x\(u,v) \ (NotR (x).M a)[u\n>v] = NotR (x).(M[u\n>v]) a" - "(NotL .M x)[u\n>v] = (if x=u then NotL .(M[u\n>v]) v else NotL .(M[u\n>v]) x)" - "\a\(N,c);b\(M,c);b\a\ \ (AndR .M .N c)[u\n>v] = AndR .(M[u\n>v]) .(N[u\n>v]) c" - "x\(u,v,y) \ (AndL1 (x).M y)[u\n>v] = (if y=u then AndL1 (x).(M[u\n>v]) v else AndL1 (x).(M[u\n>v]) y)" - "x\(u,v,y) \ (AndL2 (x).M y)[u\n>v] = (if y=u then AndL2 (x).(M[u\n>v]) v else AndL2 (x).(M[u\n>v]) y)" - "a\b \ (OrR1 .M b)[u\n>v] = OrR1 .(M[u\n>v]) b" - "a\b \ (OrR2 .M b)[u\n>v] = OrR2 .(M[u\n>v]) b" - "\x\(u,v,N,z);y\(u,v,M,z);y\x\ \ (OrL (x).M (y).N z)[u\n>v] = +| "\a\N;x\(u,v,M)\ \ (Cut .M (x).N)[u\n>v] = Cut .(M[u\n>v]) (x).(N[u\n>v])" +| "x\(u,v) \ (NotR (x).M a)[u\n>v] = NotR (x).(M[u\n>v]) a" +| "(NotL .M x)[u\n>v] = (if x=u then NotL .(M[u\n>v]) v else NotL .(M[u\n>v]) x)" +| "\a\(N,c);b\(M,c);b\a\ \ (AndR .M .N c)[u\n>v] = AndR .(M[u\n>v]) .(N[u\n>v]) c" +| "x\(u,v,y) \ (AndL1 (x).M y)[u\n>v] = (if y=u then AndL1 (x).(M[u\n>v]) v else AndL1 (x).(M[u\n>v]) y)" +| "x\(u,v,y) \ (AndL2 (x).M y)[u\n>v] = (if y=u then AndL2 (x).(M[u\n>v]) v else AndL2 (x).(M[u\n>v]) y)" +| "a\b \ (OrR1 .M b)[u\n>v] = OrR1 .(M[u\n>v]) b" +| "a\b \ (OrR2 .M b)[u\n>v] = OrR2 .(M[u\n>v]) b" +| "\x\(u,v,N,z);y\(u,v,M,z);y\x\ \ (OrL (x).M (y).N z)[u\n>v] = (if z=u then OrL (x).(M[u\n>v]) (y).(N[u\n>v]) v else OrL (x).(M[u\n>v]) (y).(N[u\n>v]) z)" - "\a\b; x\(u,v)\ \ (ImpR (x)..M b)[u\n>v] = ImpR (x)..(M[u\n>v]) b" - "\a\N;x\(u,v,M,y)\ \ (ImpL .M (x).N y)[u\n>v] = +| "\a\b; x\(u,v)\ \ (ImpR (x)..M b)[u\n>v] = ImpR (x)..(M[u\n>v]) b" +| "\a\N;x\(u,v,M,y)\ \ (ImpL .M (x).N y)[u\n>v] = (if y=u then ImpL .(M[u\n>v]) (x).(N[u\n>v]) v else ImpL .(M[u\n>v]) (x).(N[u\n>v]) y)" apply(finite_guess)+ apply(rule TrueI)+ @@ -766,32 +770,30 @@ apply(simp add: fin_supp) done -consts +nominal_primrec (freshness_context: "(y::name,c::coname,P::trm)") substn :: "trm \ name \ coname \ trm \ trm" ("_{_:=<_>._}" [100,100,100,100] 100) - substc :: "trm \ coname \ name \ trm \ trm" ("_{_:=(_)._}" [100,100,100,100] 100) - -nominal_primrec (freshness_context: "(y::name,c::coname,P::trm)") +where "(Ax x a){y:=.P} = (if x=y then Cut .P (y).Ax y a else Ax x a)" - "\a\(c,P,N);x\(y,P,M)\ \ (Cut .M (x).N){y:=.P} = +| "\a\(c,P,N);x\(y,P,M)\ \ (Cut .M (x).N){y:=.P} = (if M=Ax y a then Cut .P (x).(N{y:=.P}) else Cut .(M{y:=.P}) (x).(N{y:=.P}))" - "x\(y,P) \ (NotR (x).M a){y:=.P} = NotR (x).(M{y:=.P}) a" - "a\(c,P) \ (NotL .M x){y:=.P} = +| "x\(y,P) \ (NotR (x).M a){y:=.P} = NotR (x).(M{y:=.P}) a" +| "a\(c,P) \ (NotL .M x){y:=.P} = (if x=y then fresh_fun (\x'. Cut .P (x').NotL .(M{y:=.P}) x') else NotL .(M{y:=.P}) x)" - "\a\(c,P,N,d);b\(c,P,M,d);b\a\ \ +| "\a\(c,P,N,d);b\(c,P,M,d);b\a\ \ (AndR .M .N d){y:=.P} = AndR .(M{y:=.P}) .(N{y:=.P}) d" - "x\(y,P,z) \ (AndL1 (x).M z){y:=.P} = +| "x\(y,P,z) \ (AndL1 (x).M z){y:=.P} = (if z=y then fresh_fun (\z'. Cut .P (z').AndL1 (x).(M{y:=.P}) z') else AndL1 (x).(M{y:=.P}) z)" - "x\(y,P,z) \ (AndL2 (x).M z){y:=.P} = +| "x\(y,P,z) \ (AndL2 (x).M z){y:=.P} = (if z=y then fresh_fun (\z'. Cut .P (z').AndL2 (x).(M{y:=.P}) z') else AndL2 (x).(M{y:=.P}) z)" - "a\(c,P,b) \ (OrR1 .M b){y:=.P} = OrR1 .(M{y:=.P}) b" - "a\(c,P,b) \ (OrR2 .M b){y:=.P} = OrR2 .(M{y:=.P}) b" - "\x\(y,N,P,z);u\(y,M,P,z);x\u\ \ (OrL (x).M (u).N z){y:=.P} = +| "a\(c,P,b) \ (OrR1 .M b){y:=.P} = OrR1 .(M{y:=.P}) b" +| "a\(c,P,b) \ (OrR2 .M b){y:=.P} = OrR2 .(M{y:=.P}) b" +| "\x\(y,N,P,z);u\(y,M,P,z);x\u\ \ (OrL (x).M (u).N z){y:=.P} = (if z=y then fresh_fun (\z'. Cut .P (z').OrL (x).(M{y:=.P}) (u).(N{y:=.P}) z') else OrL (x).(M{y:=.P}) (u).(N{y:=.P}) z)" - "\a\(b,c,P); x\(y,P)\ \ (ImpR (x)..M b){y:=.P} = ImpR (x)..(M{y:=.P}) b" - "\a\(N,c,P);x\(y,P,M,z)\ \ (ImpL .M (x).N z){y:=.P} = +| "\a\(b,c,P); x\(y,P)\ \ (ImpR (x)..M b){y:=.P} = ImpR (x)..(M{y:=.P}) b" +| "\a\(N,c,P);x\(y,P,M,z)\ \ (ImpL .M (x).N z){y:=.P} = (if y=z then fresh_fun (\z'. Cut .P (z').ImpL .(M{y:=.P}) (x).(N{y:=.P}) z') else ImpL .(M{y:=.P}) (x).(N{y:=.P}) z)" apply(finite_guess)+ @@ -842,27 +844,29 @@ done nominal_primrec (freshness_context: "(d::name,z::coname,P::trm)") + substc :: "trm \ coname \ name \ trm \ trm" ("_{_:=(_)._}" [100,100,100,100] 100) +where "(Ax x a){d:=(z).P} = (if d=a then Cut .(Ax x a) (z).P else Ax x a)" - "\a\(d,P,N);x\(z,P,M)\ \ (Cut .M (x).N){d:=(z).P} = +| "\a\(d,P,N);x\(z,P,M)\ \ (Cut .M (x).N){d:=(z).P} = (if N=Ax x d then Cut .(M{d:=(z).P}) (z).P else Cut .(M{d:=(z).P}) (x).(N{d:=(z).P}))" - "x\(z,P) \ (NotR (x).M a){d:=(z).P} = +| "x\(z,P) \ (NotR (x).M a){d:=(z).P} = (if d=a then fresh_fun (\a'. Cut .NotR (x).(M{d:=(z).P}) a' (z).P) else NotR (x).(M{d:=(z).P}) a)" - "a\(d,P) \ (NotL .M x){d:=(z).P} = NotL .(M{d:=(z).P}) x" - "\a\(P,c,N,d);b\(P,c,M,d);b\a\ \ (AndR .M .N c){d:=(z).P} = +| "a\(d,P) \ (NotL .M x){d:=(z).P} = NotL .(M{d:=(z).P}) x" +| "\a\(P,c,N,d);b\(P,c,M,d);b\a\ \ (AndR .M .N c){d:=(z).P} = (if d=c then fresh_fun (\a'. Cut .(AndR .(M{d:=(z).P}) .(N{d:=(z).P}) a') (z).P) else AndR .(M{d:=(z).P}) .(N{d:=(z).P}) c)" - "x\(y,z,P) \ (AndL1 (x).M y){d:=(z).P} = AndL1 (x).(M{d:=(z).P}) y" - "x\(y,P,z) \ (AndL2 (x).M y){d:=(z).P} = AndL2 (x).(M{d:=(z).P}) y" - "a\(d,P,b) \ (OrR1 .M b){d:=(z).P} = +| "x\(y,z,P) \ (AndL1 (x).M y){d:=(z).P} = AndL1 (x).(M{d:=(z).P}) y" +| "x\(y,P,z) \ (AndL2 (x).M y){d:=(z).P} = AndL2 (x).(M{d:=(z).P}) y" +| "a\(d,P,b) \ (OrR1 .M b){d:=(z).P} = (if d=b then fresh_fun (\a'. Cut .OrR1 .(M{d:=(z).P}) a' (z).P) else OrR1 .(M{d:=(z).P}) b)" - "a\(d,P,b) \ (OrR2 .M b){d:=(z).P} = +| "a\(d,P,b) \ (OrR2 .M b){d:=(z).P} = (if d=b then fresh_fun (\a'. Cut .OrR2 .(M{d:=(z).P}) a' (z).P) else OrR2 .(M{d:=(z).P}) b)" - "\x\(N,z,P,u);y\(M,z,P,u);x\y\ \ (OrL (x).M (y).N u){d:=(z).P} = +| "\x\(N,z,P,u);y\(M,z,P,u);x\y\ \ (OrL (x).M (y).N u){d:=(z).P} = OrL (x).(M{d:=(z).P}) (y).(N{d:=(z).P}) u" - "\a\(b,d,P); x\(z,P)\ \ (ImpR (x)..M b){d:=(z).P} = +| "\a\(b,d,P); x\(z,P)\ \ (ImpR (x)..M b){d:=(z).P} = (if d=b then fresh_fun (\a'. Cut .ImpR (x)..(M{d:=(z).P}) a' (z).P) else ImpR (x)..(M{d:=(z).P}) b)" - "\a\(N,d,P);x\(y,z,P,M)\ \ (ImpL .M (x).N y){d:=(z).P} = +| "\a\(N,d,P);x\(y,z,P,M)\ \ (ImpL .M (x).N y){d:=(z).P} = ImpL .(M{d:=(z).P}) (x).(N{d:=(z).P}) y" apply(finite_guess)+ apply(rule TrueI)+ @@ -10305,11 +10309,10 @@ lemma BINDINGc_decreasing: shows "X\Y \ BINDINGc B Y \ BINDINGc B X" by (simp add: BINDINGc_def) (blast) - -consts - NOTRIGHT::"ty \ ntrm set \ ctrm set" nominal_primrec + NOTRIGHT :: "ty \ ntrm set \ ctrm set" +where "NOTRIGHT (NOT B) X = { :NotR (x).M a | a x M. fic (NotR (x).M a) a \ (x):M \ X }" apply(rule TrueI)+ done @@ -10365,11 +10368,10 @@ apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) apply(simp add: swap_simps) done - -consts - NOTLEFT::"ty \ ctrm set \ ntrm set" nominal_primrec + NOTLEFT :: "ty \ ctrm set \ ntrm set" +where "NOTLEFT (NOT B) X = { (x):NotL .M x | a x M. fin (NotL .M x) x \ :M \ X }" apply(rule TrueI)+ done @@ -10425,11 +10427,10 @@ apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) apply(simp add: swap_simps) done - -consts - ANDRIGHT::"ty \ ctrm set \ ctrm set \ ctrm set" nominal_primrec + ANDRIGHT :: "ty \ ctrm set \ ctrm set \ ctrm set" +where "ANDRIGHT (B AND C) X Y = { :AndR .M .N c | c a b M N. fic (AndR .M .N c) c \ :M \ X \ :N \ Y }" apply(rule TrueI)+ @@ -10505,10 +10506,9 @@ apply(simp) done -consts - ANDLEFT1::"ty \ ntrm set \ ntrm set" - nominal_primrec + ANDLEFT1 :: "ty \ ntrm set \ ntrm set" +where "ANDLEFT1 (B AND C) X = { (y):AndL1 (x).M y | x y M. fin (AndL1 (x).M y) y \ (x):M \ X }" apply(rule TrueI)+ done @@ -10565,10 +10565,9 @@ apply(simp add: swap_simps) done -consts - ANDLEFT2::"ty \ ntrm set \ ntrm set" - nominal_primrec + ANDLEFT2 :: "ty \ ntrm set \ ntrm set" +where "ANDLEFT2 (B AND C) X = { (y):AndL2 (x).M y | x y M. fin (AndL2 (x).M y) y \ (x):M \ X }" apply(rule TrueI)+ done @@ -10625,10 +10624,9 @@ apply(simp add: swap_simps) done -consts - ORLEFT::"ty \ ntrm set \ ntrm set \ ntrm set" - nominal_primrec + ORLEFT :: "ty \ ntrm set \ ntrm set \ ntrm set" +where "ORLEFT (B OR C) X Y = { (z):OrL (x).M (y).N z | x y z M N. fin (OrL (x).M (y).N z) z \ (x):M \ X \ (y):N \ Y }" apply(rule TrueI)+ @@ -10704,10 +10702,9 @@ apply(simp add: swap_simps) done -consts - ORRIGHT1::"ty \ ctrm set \ ctrm set" - nominal_primrec + ORRIGHT1 :: "ty \ ctrm set \ ctrm set" +where "ORRIGHT1 (B OR C) X = { :OrR1 .M b | a b M. fic (OrR1 .M b) b \ :M \ X }" apply(rule TrueI)+ done @@ -10764,10 +10761,9 @@ apply(simp) done -consts - ORRIGHT2::"ty \ ctrm set \ ctrm set" - nominal_primrec + ORRIGHT2 :: "ty \ ctrm set \ ctrm set" +where "ORRIGHT2 (B OR C) X = { :OrR2 .M b | a b M. fic (OrR2 .M b) b \ :M \ X }" apply(rule TrueI)+ done @@ -10824,10 +10820,9 @@ apply(simp) done -consts - IMPRIGHT::"ty \ ntrm set \ ctrm set \ ntrm set \ ctrm set \ ctrm set" - nominal_primrec + IMPRIGHT :: "ty \ ntrm set \ ctrm set \ ntrm set \ ctrm set \ ctrm set" +where "IMPRIGHT (B IMP C) X Y Z U= { :ImpR (x)..M b | x a b M. fic (ImpR (x)..M b) b \ (\z P. x\(z,P) \ (z):P \ Z \ (x):(M{a:=(z).P}) \ X) @@ -10954,10 +10949,9 @@ apply(perm_simp add: nsubst_eqvt fresh_right) done -consts - IMPLEFT::"ty \ ctrm set \ ntrm set \ ntrm set" - nominal_primrec + IMPLEFT :: "ty \ ctrm set \ ntrm set \ ntrm set" +where "IMPLEFT (B IMP C) X Y = { (y):ImpL .M (x).N y | x a y M N. fin (ImpL .M (x).N y) y \ :M \ X \ (x):N \ Y }" apply(rule TrueI)+ @@ -17800,23 +17794,21 @@ apply(auto) done -consts +nominal_primrec (freshness_context: "\n::(name\coname\trm)") stn :: "trm\(name\coname\trm) list\trm" - stc :: "trm\(coname\name\trm) list\trm" - -nominal_primrec (freshness_context: "\n::(name\coname\trm)") +where "stn (Ax x a) \n = lookupc x a \n" - "\a\(N,\n);x\(M,\n)\ \ stn (Cut .M (x).N) \n = (Cut .M (x).N)" - "x\\n \ stn (NotR (x).M a) \n = (NotR (x).M a)" - "a\\n \stn (NotL .M x) \n = (NotL .M x)" - "\a\(N,d,b,\n);b\(M,d,a,\n)\ \ stn (AndR .M .N d) \n = (AndR .M .N d)" - "x\(z,\n) \ stn (AndL1 (x).M z) \n = (AndL1 (x).M z)" - "x\(z,\n) \ stn (AndL2 (x).M z) \n = (AndL2 (x).M z)" - "a\(b,\n) \ stn (OrR1 .M b) \n = (OrR1 .M b)" - "a\(b,\n) \ stn (OrR2 .M b) \n = (OrR2 .M b)" - "\x\(N,z,u,\n);u\(M,z,x,\n)\ \ stn (OrL (x).M (u).N z) \n = (OrL (x).M (u).N z)" - "\a\(b,\n);x\\n\ \ stn (ImpR (x)..M b) \n = (ImpR (x)..M b)" - "\a\(N,\n);x\(M,z,\n)\ \ stn (ImpL .M (x).N z) \n = (ImpL .M (x).N z)" +| "\a\(N,\n);x\(M,\n)\ \ stn (Cut .M (x).N) \n = (Cut .M (x).N)" +| "x\\n \ stn (NotR (x).M a) \n = (NotR (x).M a)" +| "a\\n \stn (NotL .M x) \n = (NotL .M x)" +| "\a\(N,d,b,\n);b\(M,d,a,\n)\ \ stn (AndR .M .N d) \n = (AndR .M .N d)" +| "x\(z,\n) \ stn (AndL1 (x).M z) \n = (AndL1 (x).M z)" +| "x\(z,\n) \ stn (AndL2 (x).M z) \n = (AndL2 (x).M z)" +| "a\(b,\n) \ stn (OrR1 .M b) \n = (OrR1 .M b)" +| "a\(b,\n) \ stn (OrR2 .M b) \n = (OrR2 .M b)" +| "\x\(N,z,u,\n);u\(M,z,x,\n)\ \ stn (OrL (x).M (u).N z) \n = (OrL (x).M (u).N z)" +| "\a\(b,\n);x\\n\ \ stn (ImpR (x)..M b) \n = (ImpR (x)..M b)" +| "\a\(N,\n);x\(M,z,\n)\ \ stn (ImpL .M (x).N z) \n = (ImpL .M (x).N z)" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh abs_supp fin_supp)+ @@ -17824,18 +17816,20 @@ done nominal_primrec (freshness_context: "\c::(coname\name\trm)") + stc :: "trm\(coname\name\trm) list\trm" +where "stc (Ax x a) \c = lookupd x a \c" - "\a\(N,\c);x\(M,\c)\ \ stc (Cut .M (x).N) \c = (Cut .M (x).N)" - "x\\c \ stc (NotR (x).M a) \c = (NotR (x).M a)" - "a\\c \ stc (NotL .M x) \c = (NotL .M x)" - "\a\(N,d,b,\c);b\(M,d,a,\c)\ \ stc (AndR .M .N d) \c = (AndR .M .N d)" - "x\(z,\c) \ stc (AndL1 (x).M z) \c = (AndL1 (x).M z)" - "x\(z,\c) \ stc (AndL2 (x).M z) \c = (AndL2 (x).M z)" - "a\(b,\c) \ stc (OrR1 .M b) \c = (OrR1 .M b)" - "a\(b,\c) \ stc (OrR2 .M b) \c = (OrR2 .M b)" - "\x\(N,z,u,\c);u\(M,z,x,\c)\ \ stc (OrL (x).M (u).N z) \c = (OrL (x).M (u).N z)" - "\a\(b,\c);x\\c\ \ stc (ImpR (x)..M b) \c = (ImpR (x)..M b)" - "\a\(N,\c);x\(M,z,\c)\ \ stc (ImpL .M (x).N z) \c = (ImpL .M (x).N z)" +| "\a\(N,\c);x\(M,\c)\ \ stc (Cut .M (x).N) \c = (Cut .M (x).N)" +| "x\\c \ stc (NotR (x).M a) \c = (NotR (x).M a)" +| "a\\c \ stc (NotL .M x) \c = (NotL .M x)" +| "\a\(N,d,b,\c);b\(M,d,a,\c)\ \ stc (AndR .M .N d) \c = (AndR .M .N d)" +| "x\(z,\c) \ stc (AndL1 (x).M z) \c = (AndL1 (x).M z)" +| "x\(z,\c) \ stc (AndL2 (x).M z) \c = (AndL2 (x).M z)" +| "a\(b,\c) \ stc (OrR1 .M b) \c = (OrR1 .M b)" +| "a\(b,\c) \ stc (OrR2 .M b) \c = (OrR2 .M b)" +| "\x\(N,z,u,\c);u\(M,z,x,\c)\ \ stc (OrL (x).M (u).N z) \c = (OrL (x).M (u).N z)" +| "\a\(b,\c);x\\c\ \ stc (ImpR (x)..M b) \c = (ImpR (x)..M b)" +| "\a\(N,\c);x\(M,z,\c)\ \ stc (ImpL .M (x).N z) \c = (ImpL .M (x).N z)" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh abs_supp fin_supp)+ @@ -17926,51 +17920,50 @@ apply(perm_simp) done -consts +nominal_primrec (freshness_context: "(\n::(name\coname\trm) list,\c::(coname\name\trm) list)") psubst :: "(name\coname\trm) list\(coname\name\trm) list\trm\trm" ("_,_<_>" [100,100,100] 100) - -nominal_primrec (freshness_context: "(\n::(name\coname\trm) list,\c::(coname\name\trm) list)") +where "\n,\c = lookup x a \n \c" - "\a\(N,\n,\c);x\(M,\n,\c)\ \ \n,\c.M (x).N> = +| "\a\(N,\n,\c);x\(M,\n,\c)\ \ \n,\c.M (x).N> = Cut .(if \x. M=Ax x a then stn M \n else \n,\c) (x).(if \a. N=Ax x a then stc N \c else \n,\c)" - "x\(\n,\c) \ \n,\c = +| "x\(\n,\c) \ \n,\c = (case (findc \c a) of Some (u,P) \ fresh_fun (\a'. Cut .NotR (x).(\n,\c) a' (u).P) | None \ NotR (x).(\n,\c) a)" - "a\(\n,\c) \ \n,\c.M x> = +| "a\(\n,\c) \ \n,\c.M x> = (case (findn \n x) of Some (c,P) \ fresh_fun (\x'. Cut .P (x').(NotL .(\n,\c) x')) | None \ NotL .(\n,\c) x)" - "\a\(N,c,\n,\c);b\(M,c,\n,\c);b\a\ \ (\n,\c.M .N c>) = +| "\a\(N,c,\n,\c);b\(M,c,\n,\c);b\a\ \ (\n,\c.M .N c>) = (case (findc \c c) of Some (x,P) \ fresh_fun (\a'. Cut .(AndR .(\n,\c) .(\n,\c) a') (x).P) | None \ AndR .(\n,\c) .(\n,\c) c)" - "x\(z,\n,\c) \ (\n,\c) = +| "x\(z,\n,\c) \ (\n,\c) = (case (findn \n z) of Some (c,P) \ fresh_fun (\z'. Cut .P (z').AndL1 (x).(\n,\c) z') | None \ AndL1 (x).(\n,\c) z)" - "x\(z,\n,\c) \ (\n,\c) = +| "x\(z,\n,\c) \ (\n,\c) = (case (findn \n z) of Some (c,P) \ fresh_fun (\z'. Cut .P (z').AndL2 (x).(\n,\c) z') | None \ AndL2 (x).(\n,\c) z)" - "\x\(N,z,\n,\c);u\(M,z,\n,\c);x\u\ \ (\n,\c) = +| "\x\(N,z,\n,\c);u\(M,z,\n,\c);x\u\ \ (\n,\c) = (case (findn \n z) of Some (c,P) \ fresh_fun (\z'. Cut .P (z').OrL (x).(\n,\c) (u).(\n,\c) z') | None \ OrL (x).(\n,\c) (u).(\n,\c) z)" - "a\(b,\n,\c) \ (\n,\c.M b>) = +| "a\(b,\n,\c) \ (\n,\c.M b>) = (case (findc \c b) of Some (x,P) \ fresh_fun (\a'. Cut .OrR1 .(\n,\c) a' (x).P) | None \ OrR1 .(\n,\c) b)" - "a\(b,\n,\c) \ (\n,\c.M b>) = +| "a\(b,\n,\c) \ (\n,\c.M b>) = (case (findc \c b) of Some (x,P) \ fresh_fun (\a'. Cut .OrR2 .(\n,\c) a' (x).P) | None \ OrR2 .(\n,\c) b)" - "\a\(b,\n,\c); x\(\n,\c)\ \ (\n,\c.M b>) = +| "\a\(b,\n,\c); x\(\n,\c)\ \ (\n,\c.M b>) = (case (findc \c b) of Some (z,P) \ fresh_fun (\a'. Cut .ImpR (x)..(\n,\c) a' (z).P) | None \ ImpR (x)..(\n,\c) b)" - "\a\(N,\n,\c); x\(z,M,\n,\c)\ \ (\n,\c.M (x).N z>) = +| "\a\(N,\n,\c); x\(z,M,\n,\c)\ \ (\n,\c.M (x).N z>) = (case (findn \n z) of Some (c,P) \ fresh_fun (\z'. Cut .P (z').ImpL .(\n,\c) (x).(\n,\c) z') | None \ ImpL .(\n,\c) (x).(\n,\c) z)" diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/Examples/Compile.thy --- a/src/HOL/Nominal/Examples/Compile.thy Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/Examples/Compile.thy Sat Dec 13 17:13:09 2008 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - (* The definitions for a challenge suggested by Adam Chlipala *) theory Compile @@ -92,20 +90,24 @@ text {* capture-avoiding substitution *} -consts - subst :: "'a \ name \ 'a \ 'a" ("_[_::=_]" [100,100,100] 100) +class subst = + fixes subst :: "'a \ name \ 'a \ 'a" ("_[_::=_]" [100,100,100] 100) -nominal_primrec +instantiation trm :: subst +begin + +nominal_primrec subst_trm +where "(Var x)[y::=t'] = (if x=y then t' else (Var x))" - "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])" - "\x\y; x\t'\ \ (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])" - "(Const n)[y::=t'] = Const n" - "(Pr e1 e2)[y::=t'] = Pr (e1[y::=t']) (e2[y::=t'])" - "(Fst e)[y::=t'] = Fst (e[y::=t'])" - "(Snd e)[y::=t'] = Snd (e[y::=t'])" - "(InL e)[y::=t'] = InL (e[y::=t'])" - "(InR e)[y::=t'] = InR (e[y::=t'])" - "\z\x; x\y; x\e; x\e2; z\y; z\e; z\e1; x\t'; z\t'\ \ +| "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])" +| "\x\y; x\t'\ \ (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])" +| "(Const n)[y::=t'] = Const n" +| "(Pr e1 e2)[y::=t'] = Pr (e1[y::=t']) (e2[y::=t'])" +| "(Fst e)[y::=t'] = Fst (e[y::=t'])" +| "(Snd e)[y::=t'] = Snd (e[y::=t'])" +| "(InL e)[y::=t'] = InL (e[y::=t'])" +| "(InR e)[y::=t'] = InR (e[y::=t'])" +| "\z\x; x\y; x\e; x\e2; z\y; z\e; z\e1; x\t'; z\t'\ \ (Case e of inl x \ e1 | inr z \ e2)[y::=t'] = (Case (e[y::=t']) of inl x \ (e1[y::=t']) | inr z \ (e2[y::=t']))" apply(finite_guess)+ @@ -114,23 +116,35 @@ apply(fresh_guess)+ done -nominal_primrec (Isubst) +instance .. + +end + +instantiation trmI :: subst +begin + +nominal_primrec subst_trmI +where "(IVar x)[y::=t'] = (if x=y then t' else (IVar x))" - "(IApp t1 t2)[y::=t'] = IApp (t1[y::=t']) (t2[y::=t'])" - "\x\y; x\t'\ \ (ILam [x].t)[y::=t'] = ILam [x].(t[y::=t'])" - "(INat n)[y::=t'] = INat n" - "(IUnit)[y::=t'] = IUnit" - "(ISucc e)[y::=t'] = ISucc (e[y::=t'])" - "(IAss e1 e2)[y::=t'] = IAss (e1[y::=t']) (e2[y::=t'])" - "(IRef e)[y::=t'] = IRef (e[y::=t'])" - "(ISeq e1 e2)[y::=t'] = ISeq (e1[y::=t']) (e2[y::=t'])" - "(Iif e e1 e2)[y::=t'] = Iif (e[y::=t']) (e1[y::=t']) (e2[y::=t'])" +| "(IApp t1 t2)[y::=t'] = IApp (t1[y::=t']) (t2[y::=t'])" +| "\x\y; x\t'\ \ (ILam [x].t)[y::=t'] = ILam [x].(t[y::=t'])" +| "(INat n)[y::=t'] = INat n" +| "(IUnit)[y::=t'] = IUnit" +| "(ISucc e)[y::=t'] = ISucc (e[y::=t'])" +| "(IAss e1 e2)[y::=t'] = IAss (e1[y::=t']) (e2[y::=t'])" +| "(IRef e)[y::=t'] = IRef (e[y::=t'])" +| "(ISeq e1 e2)[y::=t'] = ISeq (e1[y::=t']) (e2[y::=t'])" +| "(Iif e e1 e2)[y::=t'] = Iif (e[y::=t']) (e1[y::=t']) (e2[y::=t'])" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh)+ apply(fresh_guess)+ done +instance .. + +end + lemma Isubst_eqvt[eqvt]: fixes pi::"name prm" and t1::"trmI" @@ -138,7 +152,7 @@ and x::"name" shows "pi\(t1[x::=t2]) = ((pi\t1)[(pi\x)::=(pi\t2)])" apply (nominal_induct t1 avoiding: x t2 rule: trmI.strong_induct) - apply (simp_all add: Isubst.simps eqvts fresh_bij) + apply (simp_all add: subst_trmI.simps eqvts fresh_bij) done lemma Isubst_supp: @@ -147,7 +161,7 @@ and x::"name" shows "((supp (t1[x::=t2]))::name set) \ (supp t2)\((supp t1)-{x})" apply (nominal_induct t1 avoiding: x t2 rule: trmI.strong_induct) - apply (auto simp add: Isubst.simps trmI.supp supp_atm abs_supp supp_nat) + apply (auto simp add: subst_trmI.simps trmI.supp supp_atm abs_supp supp_nat) apply blast+ done @@ -198,29 +212,29 @@ text {* Translation functions *} -consts trans :: "trm \ trmI" - nominal_primrec + trans :: "trm \ trmI" +where "trans (Var x) = (IVar x)" - "trans (App e1 e2) = IApp (trans e1) (trans e2)" - "trans (Lam [x].e) = ILam [x].(trans e)" - "trans (Const n) = INat n" - "trans (Pr e1 e2) = +| "trans (App e1 e2) = IApp (trans e1) (trans e2)" +| "trans (Lam [x].e) = ILam [x].(trans e)" +| "trans (Const n) = INat n" +| "trans (Pr e1 e2) = (let limit = IRef(INat 0) in let v1 = (trans e1) in let v2 = (trans e2) in (((ISucc limit)\v1);;(ISucc(ISucc limit)\v2));;(INat 0 \ ISucc(ISucc(limit))))" - "trans (Fst e) = IRef (ISucc (trans e))" - "trans (Snd e) = IRef (ISucc (ISucc (trans e)))" - "trans (InL e) = +| "trans (Fst e) = IRef (ISucc (trans e))" +| "trans (Snd e) = IRef (ISucc (ISucc (trans e)))" +| "trans (InL e) = (let limit = IRef(INat 0) in let v = (trans e) in (((ISucc limit)\INat(0));;(ISucc(ISucc limit)\v));;(INat 0 \ ISucc(ISucc(limit))))" - "trans (InR e) = +| "trans (InR e) = (let limit = IRef(INat 0) in let v = (trans e) in (((ISucc limit)\INat(1));;(ISucc(ISucc limit)\v));;(INat 0 \ ISucc(ISucc(limit))))" - "\x2\x1; x1\e; x1\e2; x2\e; x2\e1\ \ +| "\x2\x1; x1\e; x1\e2; x2\e; x2\e1\ \ trans (Case e of inl x1 \ e1 | inr x2 \ e2) = (let v = (trans e) in let v1 = (trans e1) in @@ -232,11 +246,11 @@ apply(fresh_guess add: Let_def)+ done -consts trans_type :: "ty \ tyI" - nominal_primrec + trans_type :: "ty \ tyI" +where "trans_type (Data \) = DataI(NatI)" - "trans_type (\1\\2) = (trans_type \1)\(trans_type \2)" +| "trans_type (\1\\2) = (trans_type \1)\(trans_type \2)" by (rule TrueI)+ end \ No newline at end of file diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/Examples/Contexts.thy --- a/src/HOL/Nominal/Examples/Contexts.thy Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/Examples/Contexts.thy Sat Dec 13 17:13:09 2008 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory Contexts imports "../Nominal" begin @@ -42,12 +40,12 @@ text {* Capture-Avoiding Substitution *} -consts subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [100,100,100] 100) - nominal_primrec + subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [100,100,100] 100) +where "(Var x)[y::=s] = (if x=y then s else (Var x))" - "(App t\<^isub>1 t\<^isub>2)[y::=s] = App (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])" - "x\(y,s) \ (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])" +| "(App t\<^isub>1 t\<^isub>2)[y::=s] = App (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])" +| "x\(y,s) \ (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh) @@ -59,14 +57,13 @@ This operation is possibly capturing. *} -consts +nominal_primrec filling :: "ctx \ lam \ lam" ("_\_\" [100,100] 100) - -nominal_primrec +where "\\t\ = t" - "(CAppL E t')\t\ = App (E\t\) t'" - "(CAppR t' E)\t\ = App t' (E\t\)" - "(CLam [x].E)\t\ = Lam [x].(E\t\)" +| "(CAppL E t')\t\ = App (E\t\) t'" +| "(CAppR t' E)\t\ = App t' (E\t\)" +| "(CLam [x].E)\t\ = Lam [x].(E\t\)" by (rule TrueI)+ text {* @@ -81,14 +78,13 @@ text {* The composition of two contexts. *} -consts +nominal_primrec ctx_compose :: "ctx \ ctx \ ctx" ("_ \ _" [100,100] 100) - -nominal_primrec +where "\ \ E' = E'" - "(CAppL E t') \ E' = CAppL (E \ E') t'" - "(CAppR t' E) \ E' = CAppR t' (E \ E')" - "(CLam [x].E) \ E' = CLam [x].(E \ E')" +| "(CAppL E t') \ E' = CAppL (E \ E') t'" +| "(CAppR t' E) \ E' = CAppR t' (E \ E')" +| "(CLam [x].E) \ E' = CLam [x].(E \ E')" by (rule TrueI)+ lemma ctx_compose: diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/Examples/Crary.thy --- a/src/HOL/Nominal/Examples/Crary.thy Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/Examples/Crary.thy Sat Dec 13 17:13:09 2008 +0100 @@ -1,4 +1,3 @@ -(* "$Id$" *) (* *) (* Formalisation of the chapter on Logical Relations *) (* and a Case Study in Equivalence Checking *) @@ -47,14 +46,20 @@ shows "(\ T\<^isub>1 T\<^isub>2. T=T\<^isub>1\T\<^isub>2) \ T=TUnit \ T=TBase" by (induct T rule:ty.induct) (auto) -instance ty :: size .. +instantiation ty :: size +begin -nominal_primrec +nominal_primrec size_ty +where "size (TBase) = 1" - "size (TUnit) = 1" - "size (T\<^isub>1\T\<^isub>2) = size T\<^isub>1 + size T\<^isub>2" +| "size (TUnit) = 1" +| "size (T\<^isub>1\T\<^isub>2) = size T\<^isub>1 + size T\<^isub>2" by (rule TrueI)+ +instance .. + +end + lemma ty_size_greater_zero[simp]: fixes T::"ty" shows "size T > 0" @@ -87,16 +92,15 @@ using a by (induct rule: lookup.induct) (auto simp add: fresh_list_cons fresh_prod fresh_atm) - -consts - psubst :: "Subst \ trm \ trm" ("_<_>" [100,100] 130) nominal_primrec + psubst :: "Subst \ trm \ trm" ("_<_>" [100,100] 130) +where "\<(Var x)> = (lookup \ x)" - "\<(App t\<^isub>1 t\<^isub>2)> = App \1> \2>" - "x\\ \ \<(Lam [x].t)> = Lam [x].(\)" - "\<(Const n)> = Const n" - "\<(Unit)> = Unit" +| "\<(App t\<^isub>1 t\<^isub>2)> = App \1> \2>" +| "x\\ \ \<(Lam [x].t)> = Lam [x].(\)" +| "\<(Const n)> = Const n" +| "\<(Unit)> = Unit" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh)+ diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/Examples/Fsub.thy --- a/src/HOL/Nominal/Examples/Fsub.thy Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/Examples/Fsub.thy Sat Dec 13 17:13:09 2008 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - (*<*) theory Fsub imports "../Nominal" @@ -229,32 +227,26 @@ section {* Size and Capture-Avoiding Substitution for Types *} -consts size_ty :: "ty \ nat" - nominal_primrec + size_ty :: "ty \ nat" +where "size_ty (Tvar X) = 1" - "size_ty (Top) = 1" - "size_ty (T1 \ T2) = (size_ty T1) + (size_ty T2) + 1" - "X\T1 \ size_ty (\[X<:T1].T2) = (size_ty T1) + (size_ty T2) + 1" +| "size_ty (Top) = 1" +| "size_ty (T1 \ T2) = (size_ty T1) + (size_ty T2) + 1" +| "X\T1 \ size_ty (\[X<:T1].T2) = (size_ty T1) + (size_ty T2) + 1" apply (finite_guess)+ apply (rule TrueI)+ apply (simp add: fresh_nat) apply (fresh_guess)+ done -consts subst_ty :: "tyvrs \ ty \ ty \ ty" - -syntax - subst_ty_syn :: "ty \ tyvrs \ ty \ ty" ("_[_:=_]\<^isub>t\<^isub>y" [100,100,100] 100) - -translations - "T1[Y:=T2]\<^isub>t\<^isub>y" \ "subst_ty Y T2 T1" - nominal_primrec + subst_ty :: "ty \ tyvrs \ ty \ ty" ("_[_:=_]\<^isub>t\<^isub>y" [100,100,100] 100) +where "(Tvar X)[Y:=T]\<^isub>t\<^isub>y= (if X=Y then T else (Tvar X))" - "(Top)[Y:=T]\<^isub>t\<^isub>y = Top" - "(T\<^isub>1 \ T\<^isub>2)[Y:=T]\<^isub>t\<^isub>y = (T\<^isub>1[Y:=T]\<^isub>t\<^isub>y) \ (T\<^isub>2[Y:=T]\<^isub>t\<^isub>y)" - "\X\(Y,T); X\T\<^isub>1\ \ (\[X<:T\<^isub>1].T\<^isub>2)[Y:=T]\<^isub>t\<^isub>y = (\[X<:(T\<^isub>1[Y:=T]\<^isub>t\<^isub>y)].(T\<^isub>2[Y:=T]\<^isub>t\<^isub>y))" +| "(Top)[Y:=T]\<^isub>t\<^isub>y = Top" +| "(T\<^isub>1 \ T\<^isub>2)[Y:=T]\<^isub>t\<^isub>y = (T\<^isub>1[Y:=T]\<^isub>t\<^isub>y) \ (T\<^isub>2[Y:=T]\<^isub>t\<^isub>y)" +| "\X\(Y,T); X\T\<^isub>1\ \ (\[X<:T\<^isub>1].T\<^isub>2)[Y:=T]\<^isub>t\<^isub>y = (\[X<:(T\<^isub>1[Y:=T]\<^isub>t\<^isub>y)].(T\<^isub>2[Y:=T]\<^isub>t\<^isub>y))" apply (finite_guess)+ apply (rule TrueI)+ apply (simp add: abs_fresh) diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/Examples/Height.thy --- a/src/HOL/Nominal/Examples/Height.thy Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/Examples/Height.thy Sat Dec 13 17:13:09 2008 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory Height imports "../Nominal" begin @@ -17,13 +15,13 @@ | Lam "\name\lam" ("Lam [_]._" [100,100] 100) text {* Definition of the height-function on lambda-terms. *} -consts - height :: "lam \ int" nominal_primrec + height :: "lam \ int" +where "height (Var x) = 1" - "height (App t1 t2) = (max (height t1) (height t2)) + 1" - "height (Lam [a].t) = (height t) + 1" +| "height (App t1 t2) = (max (height t1) (height t2)) + 1" +| "height (Lam [a].t) = (height t) + 1" apply(finite_guess add: perm_int_def)+ apply(rule TrueI)+ apply(simp add: fresh_int) @@ -32,13 +30,12 @@ text {* Definition of capture-avoiding substitution. *} -consts +nominal_primrec subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [100,100,100] 100) - -nominal_primrec +where "(Var x)[y::=t'] = (if x=y then t' else (Var x))" - "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])" - "\x\y; x\t'\ \ (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])" +| "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])" +| "\x\y; x\t'\ \ (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh) diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/Examples/Lam_Funs.thy --- a/src/HOL/Nominal/Examples/Lam_Funs.thy Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/Examples/Lam_Funs.thy Sat Dec 13 17:13:09 2008 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory Lam_Funs imports "../Nominal" begin @@ -18,13 +16,12 @@ text {* The depth of a lambda-term. *} -consts +nominal_primrec depth :: "lam \ nat" - -nominal_primrec +where "depth (Var x) = 1" - "depth (App t1 t2) = (max (depth t1) (depth t2)) + 1" - "depth (Lam [a].t) = (depth t) + 1" +| "depth (App t1 t2) = (max (depth t1) (depth t2)) + 1" +| "depth (Lam [a].t) = (depth t) + 1" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: fresh_nat) @@ -38,13 +35,12 @@ the invariant that frees always returns a finite set of names. *} -consts +nominal_primrec (invariant: "\s::name set. finite s") frees :: "lam \ name set" - -nominal_primrec (invariant: "\s::name set. finite s") +where "frees (Var a) = {a}" - "frees (App t1 t2) = (frees t1) \ (frees t2)" - "frees (Lam [a].t) = (frees t) - {a}" +| "frees (App t1 t2) = (frees t1) \ (frees t2)" +| "frees (Lam [a].t) = (frees t) - {a}" apply(finite_guess)+ apply(simp)+ apply(simp add: fresh_def) @@ -78,14 +74,13 @@ and X::"name" shows "pi\(lookup \ X) = lookup (pi\\) (pi\X)" by (induct \) (auto simp add: eqvts) - -consts - psubst :: "(name\lam) list \ lam \ lam" ("_<_>" [95,95] 105) nominal_primrec + psubst :: "(name\lam) list \ lam \ lam" ("_<_>" [95,95] 105) +where "\<(Var x)> = (lookup \ x)" - "\<(App e\<^isub>1 e\<^isub>2)> = App (\1>) (\2>)" - "x\\ \ \<(Lam [x].e)> = Lam [x].(\)" +| "\<(App e\<^isub>1 e\<^isub>2)> = App (\1>) (\2>)" +| "x\\ \ \<(Lam [x].e)> = Lam [x].(\)" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh)+ @@ -130,26 +125,24 @@ text {* Filling a lambda-term into a context. *} -consts +nominal_primrec filling :: "clam \ lam \ lam" ("_\_\" [100,100] 100) - -nominal_primrec +where "\\t\ = t" - "(CAppL E t')\t\ = App (E\t\) t'" - "(CAppR t' E)\t\ = App t' (E\t\)" - "(CLam [x].E)\t\ = Lam [x].(E\t\)" +| "(CAppL E t')\t\ = App (E\t\) t'" +| "(CAppR t' E)\t\ = App t' (E\t\)" +| "(CLam [x].E)\t\ = Lam [x].(E\t\)" by (rule TrueI)+ text {* Composition od two contexts *} -consts +nominal_primrec clam_compose :: "clam \ clam \ clam" ("_ \ _" [100,100] 100) - -nominal_primrec +where "\ \ E' = E'" - "(CAppL E t') \ E' = CAppL (E \ E') t'" - "(CAppR t' E) \ E' = CAppR t' (E \ E')" - "(CLam [x].E) \ E' = CLam [x].(E \ E')" +| "(CAppL E t') \ E' = CAppL (E \ E') t'" +| "(CAppR t' E) \ E' = CAppR t' (E \ E')" +| "(CLam [x].E) \ E' = CLam [x].(E \ E')" by (rule TrueI)+ lemma clam_compose: diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/Examples/LocalWeakening.thy --- a/src/HOL/Nominal/Examples/LocalWeakening.thy Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/Examples/LocalWeakening.thy Sat Dec 13 17:13:09 2008 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - (* Formalisation of weakening using locally nameless *) (* terms; the nominal infrastructure can also derive *) (* strong induction principles for such representations *) @@ -29,13 +27,13 @@ by (induct t rule: llam.induct) (auto simp add: llam.inject) -consts llam_size :: "llam \ nat" - nominal_primrec - "llam_size (lPar a) = 1" - "llam_size (lVar n) = 1" - "llam_size (lApp t1 t2) = 1 + (llam_size t1) + (llam_size t2)" - "llam_size (lLam t) = 1 + (llam_size t)" + llam_size :: "llam \ nat" +where + "llam_size (lPar a) = 1" +| "llam_size (lVar n) = 1" +| "llam_size (lApp t1 t2) = 1 + (llam_size t1) + (llam_size t2)" +| "llam_size (lLam t) = 1 + (llam_size t)" by (rule TrueI)+ function diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/Examples/SN.thy --- a/src/HOL/Nominal/Examples/SN.thy Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/Examples/SN.thy Sat Dec 13 17:13:09 2008 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory SN imports Lam_Funs begin @@ -228,12 +226,11 @@ section {* Candidates *} -consts +nominal_primrec RED :: "ty \ lam set" - -nominal_primrec +where "RED (TVar X) = {t. SN(t)}" - "RED (\\\) = {t. \u. (u\RED \ \ (App t u)\RED \)}" +| "RED (\\\) = {t. \u. (u\RED \ \ (App t u)\RED \)}" by (rule TrueI)+ text {* neutral terms *} @@ -248,13 +245,12 @@ where fst[intro!]: "(App t s) \ t" -consts +nominal_primrec fst_app_aux::"lam\lam option" - -nominal_primrec +where "fst_app_aux (Var a) = None" - "fst_app_aux (App t1 t2) = Some t1" - "fst_app_aux (Lam [x].t) = None" +| "fst_app_aux (App t1 t2) = Some t1" +| "fst_app_aux (Lam [x].t) = None" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: fresh_none) diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/Examples/SOS.thy --- a/src/HOL/Nominal/Examples/SOS.thy Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/Examples/SOS.thy Sat Dec 13 17:13:09 2008 +0100 @@ -1,4 +1,3 @@ -(* "$Id$" *) (* *) (* Formalisation of some typical SOS-proofs. *) (* *) @@ -62,13 +61,12 @@ (* parallel substitution *) -consts +nominal_primrec psubst :: "(name\trm) list \ trm \ trm" ("_<_>" [95,95] 105) - -nominal_primrec +where "\<(Var x)> = (lookup \ x)" - "\<(App e\<^isub>1 e\<^isub>2)> = App (\1>) (\2>)" - "x\\ \ \<(Lam [x].e)> = Lam [x].(\)" +| "\<(App e\<^isub>1 e\<^isub>2)> = App (\1>) (\2>)" +| "x\\ \ \<(Lam [x].e)> = Lam [x].(\)" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh)+ @@ -349,12 +347,12 @@ using h by (induct) (auto) (* Valuation *) -consts - V :: "ty \ trm set" nominal_primrec + V :: "ty \ trm set" +where "V (TVar x) = {e. val e}" - "V (T\<^isub>1 \ T\<^isub>2) = {Lam [x].e | x e. \ v \ (V T\<^isub>1). \ v'. e[x::=v] \ v' \ v' \ V T\<^isub>2}" +| "V (T\<^isub>1 \ T\<^isub>2) = {Lam [x].e | x e. \ v \ (V T\<^isub>1). \ v'. e[x::=v] \ v' \ v' \ V T\<^isub>2}" by (rule TrueI)+ lemma V_eqvt: diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/Examples/Standardization.thy --- a/src/HOL/Nominal/Examples/Standardization.thy Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/Examples/Standardization.thy Sat Dec 13 17:13:09 2008 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/Nominal/Examples/Standardization.thy - ID: $Id$ Author: Stefan Berghofer and Tobias Nipkow Copyright 2005, 2008 TU Muenchen *) @@ -24,24 +23,30 @@ | App "lam" "lam" (infixl "\" 200) | Lam "\name\lam" ("Lam [_]._" [0, 10] 10) -instance lam :: size .. +instantiation lam :: size +begin -nominal_primrec +nominal_primrec size_lam +where "size (Var n) = 0" - "size (t \ u) = size t + size u + 1" - "size (Lam [x].t) = size t + 1" +| "size (t \ u) = size t + size u + 1" +| "size (Lam [x].t) = size t + 1" apply finite_guess+ apply (rule TrueI)+ apply (simp add: fresh_nat) apply fresh_guess+ done -consts subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [300, 0, 0] 300) +instance .. + +end nominal_primrec + subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [300, 0, 0] 300) +where subst_Var: "(Var x)[y::=s] = (if x=y then s else (Var x))" - subst_App: "(t\<^isub>1 \ t\<^isub>2)[y::=s] = t\<^isub>1[y::=s] \ t\<^isub>2[y::=s]" - subst_Lam: "x \ (y, s) \ (Lam [x].t)[y::=s] = (Lam [x].(t[y::=s]))" +| subst_App: "(t\<^isub>1 \ t\<^isub>2)[y::=s] = t\<^isub>1[y::=s] \ t\<^isub>2[y::=s]" +| subst_Lam: "x \ (y, s) \ (Lam [x].t)[y::=s] = (Lam [x].(t[y::=s]))" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh) diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/Examples/Type_Preservation.thy --- a/src/HOL/Nominal/Examples/Type_Preservation.thy Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/Examples/Type_Preservation.thy Sat Dec 13 17:13:09 2008 +0100 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory Type_Preservation imports Nominal begin @@ -21,13 +19,12 @@ text {* Capture-Avoiding Substitution *} -consts +nominal_primrec subst :: "lam \ name \ lam \ lam" ("_[_::=_]") - -nominal_primrec +where "(Var x)[y::=s] = (if x=y then s else (Var x))" - "(App t\<^isub>1 t\<^isub>2)[y::=s] = App (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])" - "x\(y,s) \ (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])" +| "(App t\<^isub>1 t\<^isub>2)[y::=s] = App (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])" +| "x\(y,s) \ (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh) diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/Examples/W.thy --- a/src/HOL/Nominal/Examples/W.thy Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/Examples/W.thy Sat Dec 13 17:13:09 2008 +0100 @@ -1,5 +1,3 @@ -(* "$Id$" *) - theory W imports Nominal begin @@ -50,26 +48,68 @@ Ctxt = "(var\tyS) list" text {* free type variables *} -consts - ftv :: "'a \ tvar list" + +class ftv = type + + fixes ftv :: "'a \ tvar list" + +instantiation * :: (ftv, ftv) ftv +begin + +primrec ftv_prod +where + "ftv (x::'a::ftv, y::'b::ftv) = (ftv x)@(ftv y)" + +instance .. + +end -primrec (ftv_of_prod) - "ftv (x,y) = (ftv x)@(ftv y)" +instantiation tvar :: ftv +begin + +definition + ftv_of_tvar[simp]: "ftv X \ [(X::tvar)]" -defs (overloaded) - ftv_of_tvar[simp]: "ftv X \ [(X::tvar)]" +instance .. + +end + +instantiation var :: ftv +begin + +definition ftv_of_var[simp]: "ftv (x::var) \ []" -primrec (ftv_of_list) +instance .. + +end + +instantiation list :: (ftv) ftv +begin + +primrec ftv_list +where "ftv [] = []" - "ftv (x#xs) = (ftv x)@(ftv xs)" +| "ftv (x#xs) = (ftv x)@(ftv xs)" + +instance .. + +end (* free type-variables of types *) -nominal_primrec (ftv_ty) + +instantiation ty :: ftv +begin + +nominal_primrec ftv_ty +where "ftv (TVar X) = [X]" - "ftv (T\<^isub>1\T\<^isub>2) = (ftv T\<^isub>1)@(ftv T\<^isub>2)" +| "ftv (T\<^isub>1\T\<^isub>2) = (ftv T\<^isub>1)@(ftv T\<^isub>2)" by (rule TrueI)+ +instance .. + +end + lemma ftv_ty_eqvt[eqvt]: fixes pi::"tvar prm" and T::"ty" @@ -77,9 +117,13 @@ by (nominal_induct T rule: ty.strong_induct) (perm_simp add: append_eqvt)+ -nominal_primrec (ftv_tyS) +instantiation tyS :: ftv +begin + +nominal_primrec ftv_tyS +where "ftv (Ty T) = ftv T" - "ftv (\[X].S) = (ftv S) - [X]" +| "ftv (\[X].S) = (ftv S) - [X]" apply(finite_guess add: ftv_ty_eqvt fs_tvar1)+ apply(rule TrueI)+ apply(rule difference_fresh) @@ -87,6 +131,10 @@ apply(fresh_guess add: ftv_ty_eqvt fs_tvar1)+ done +instance .. + +end + lemma ftv_tyS_eqvt[eqvt]: fixes pi::"tvar prm" and S::"tyS" @@ -140,11 +188,11 @@ types Subst = "(tvar\ty) list" -consts - psubst :: "Subst \ 'a \ 'a" ("_<_>" [100,60] 120) +class psubst = + fixes psubst :: "Subst \ 'a \ 'a" ("_<_>" [100,60] 120) abbreviation - subst :: "'a \ tvar \ ty \ 'a" ("_[_::=_]" [100,100,100] 100) + subst :: "'a::psubst \ tvar \ ty \ 'a" ("_[_::=_]" [100,100,100] 100) where "smth[X::=T] \ ([(X,T)])" @@ -159,11 +207,19 @@ shows "pi\(lookup \ X) = lookup (pi\\) (pi\X)" by (induct \) (auto simp add: eqvts) -nominal_primrec (psubst_ty) +instantiation ty :: psubst +begin + +nominal_primrec psubst_ty +where "\ = lookup \ X" - "\1 \ T\<^isub>2> = (\1>) \ (\2>)" +| "\1 \ T\<^isub>2> = (\1>) \ (\2>)" by (rule TrueI)+ +instance .. + +end + lemma psubst_ty_eqvt[eqvt]: fixes pi1::"tvar prm" and \::"Subst" diff -r 1b8e021e8c1f -r 66fe138979f4 src/HOL/Nominal/nominal_primrec.ML --- a/src/HOL/Nominal/nominal_primrec.ML Sat Dec 13 15:35:29 2008 +0100 +++ b/src/HOL/Nominal/nominal_primrec.ML Sat Dec 13 17:13:09 2008 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/Nominal/nominal_primrec.ML - ID: $Id$ Author: Stefan Berghofer, TU Muenchen and Norbert Voelker, FernUni Hagen Package for defining functions on nominal datatypes by primitive recursion. @@ -8,14 +7,10 @@ signature NOMINAL_PRIMREC = sig - val add_primrec: string -> string list option -> string option -> - ((Binding.T * string) * Attrib.src list) list -> theory -> Proof.state - val add_primrec_unchecked: string -> string list option -> string option -> - ((Binding.T * string) * Attrib.src list) list -> theory -> Proof.state - val add_primrec_i: string -> term list option -> term option -> - ((Binding.T * term) * attribute list) list -> theory -> Proof.state - val add_primrec_unchecked_i: string -> term list option -> term option -> - ((Binding.T * term) * attribute list) list -> theory -> Proof.state + val add_primrec: term list option -> term option -> + (Binding.T * typ option * mixfix) list -> + (Binding.T * typ option * mixfix) list -> + (Attrib.binding * term) list -> local_theory -> Proof.state end; structure NominalPrimrec : NOMINAL_PRIMREC = @@ -26,23 +21,31 @@ exception RecError of string; fun primrec_err s = error ("Nominal primrec definition error:\n" ^ s); -fun primrec_eq_err thy s eq = - primrec_err (s ^ "\nin\n" ^ quote (Syntax.string_of_term_global thy eq)); +fun primrec_eq_err lthy s eq = + primrec_err (s ^ "\nin\n" ^ quote (Syntax.string_of_term lthy eq)); (* preprocessing of equations *) -fun process_eqn thy eq rec_fns = +fun unquantify t = let + val (vs, Ts) = split_list (strip_qnt_vars "all" t); + val body = strip_qnt_body "all" t; + val (vs', _) = Name.variants vs (Name.make_context (fold_aterms + (fn Free (v, _) => insert (op =) v | _ => I) body [])) + in (curry subst_bounds (map2 (curry Free) vs' Ts |> rev) body) end; + +fun process_eqn lthy is_fixed spec rec_fns = + let + val eq = unquantify spec; val (lhs, rhs) = - if null (term_vars eq) then - HOLogic.dest_eq (HOLogic.dest_Trueprop (Logic.strip_imp_concl eq)) - handle TERM _ => raise RecError "not a proper equation" - else raise RecError "illegal schematic variable(s)"; + HOLogic.dest_eq (HOLogic.dest_Trueprop (Logic.strip_imp_concl eq)) + handle TERM _ => raise RecError "not a proper equation"; val (recfun, args) = strip_comb lhs; - val fnameT = dest_Const recfun handle TERM _ => - raise RecError "function is not declared as constant in theory"; + val fname = case recfun of Free (v, _) => if is_fixed v then v + else raise RecError "illegal head of function equation" + | _ => raise RecError "illegal head of function equation"; val (ls', rest) = take_prefix is_Free args; val (middle, rs') = take_suffix is_Free rest; @@ -68,26 +71,28 @@ else (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees); check_vars "extra variables on rhs: " - (map dest_Free (term_frees rhs) \\ lfrees); - case AList.lookup (op =) rec_fns fnameT of + (map dest_Free (term_frees rhs) |> subtract (op =) lfrees + |> filter_out (is_fixed o fst)); + case AList.lookup (op =) rec_fns fname of NONE => - (fnameT, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns + (fname, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns | SOME (_, rpos', eqns) => if AList.defined (op =) eqns cname then raise RecError "constructor already occurred as pattern" else if rpos <> rpos' then raise RecError "position of recursive argument inconsistent" else - AList.update (op =) (fnameT, (tname, rpos, (cname, (ls, cargs, rs, rhs, eq))::eqns)) + AList.update (op =) + (fname, (tname, rpos, (cname, (ls, cargs, rs, rhs, eq))::eqns)) rec_fns) end - handle RecError s => primrec_eq_err thy s eq; + handle RecError s => primrec_eq_err lthy s spec; val param_err = "Parameters must be the same for all recursive functions"; -fun process_fun thy descr rec_eqns (i, fnameT as (fname, _)) (fnameTs, fnss) = +fun process_fun lthy descr eqns (i, fname) (fnames, fnss) = let - val (_, (tname, _, constrs)) = List.nth (descr, i); + val (_, (tname, _, constrs)) = nth descr i; (* substitute "fname ls x rs" by "y" for (x, (_, y)) in subs *) @@ -100,16 +105,17 @@ let val (f, ts) = strip_comb t; in - if is_Const f andalso dest_Const f mem map fst rec_eqns then + if is_Free f + andalso member (fn ((v, _), (w, _)) => v = w) eqns (dest_Free f) then let - val fnameT' as (fname', _) = dest_Const f; - val (_, rpos, eqns) = the (AList.lookup (op =) rec_eqns fnameT'); - val ls = Library.take (rpos, ts); - val rest = Library.drop (rpos, ts); - val (x', rs) = (hd rest, tl rest) - handle Empty => raise RecError ("not enough arguments\ - \ in recursive application\nof function " ^ quote fname' ^ " on rhs"); - val rs' = (case eqns of + val (fname', _) = dest_Free f; + val (_, rpos, eqns') = the (AList.lookup (op =) eqns fname'); + val (ls, rs'') = chop rpos ts + val (x', rs) = case rs'' of + x' :: rs => (x', rs) + | [] => raise RecError ("not enough arguments in recursive application\n" + ^ "of function " ^ quote fname' ^ " on rhs"); + val rs' = (case eqns' of (_, (ls', _, rs', _, _)) :: _ => let val (rs1, rs2) = chop (length rs') rs in @@ -126,7 +132,7 @@ | SOME (i', y) => fs |> fold_map (subst subs) (xs @ rs') - ||> process_fun thy descr rec_eqns (i', fnameT') + ||> process_fun lthy descr eqns (i', fname') |-> (fn ts' => pair (list_comb (y, ts'))) end else @@ -138,41 +144,39 @@ (* translate rec equations into function arguments suitable for rec comb *) - fun trans eqns (cname, cargs) (fnameTs', fnss', fns) = + fun trans eqns (cname, cargs) (fnames', fnss', fns) = (case AList.lookup (op =) eqns cname of NONE => (warning ("No equation for constructor " ^ quote cname ^ "\nin definition of function " ^ quote fname); - (fnameTs', fnss', (Const (@{const_name undefined}, dummyT))::fns)) + (fnames', fnss', (Const (@{const_name undefined}, dummyT))::fns)) | SOME (ls, cargs', rs, rhs, eq) => let val recs = filter (is_rec_type o snd) (cargs' ~~ cargs); val rargs = map fst recs; - val subs = map (rpair dummyT o fst) + val subs = map (rpair dummyT o fst) (rev (rename_wrt_term rhs rargs)); - val (rhs', (fnameTs'', fnss'')) = - (subst (map (fn ((x, y), z) => - (Free x, (body_index y, Free z))) - (recs ~~ subs)) rhs (fnameTs', fnss')) - handle RecError s => primrec_eq_err thy s eq - in (fnameTs'', fnss'', + val (rhs', (fnames'', fnss'')) = subst (map2 (fn (x, y) => fn z => + (Free x, (body_index y, Free z))) recs subs) rhs (fnames', fnss') + handle RecError s => primrec_eq_err lthy s eq + in (fnames'', fnss'', (list_abs_free (cargs' @ subs, rhs'))::fns) end) - in (case AList.lookup (op =) fnameTs i of + in (case AList.lookup (op =) fnames i of NONE => - if exists (equal fnameT o snd) fnameTs then + if exists (fn (_, v) => fname = v) fnames then raise RecError ("inconsistent functions for datatype " ^ quote tname) else let - val SOME (_, _, eqns as (_, (ls, _, rs, _, _)) :: _) = - AList.lookup (op =) rec_eqns fnameT; - val (fnameTs', fnss', fns) = fold_rev (trans eqns) constrs - ((i, fnameT)::fnameTs, fnss, []) + val SOME (_, _, eqns' as (_, (ls, _, rs, _, _)) :: _) = + AList.lookup (op =) eqns fname; + val (fnames', fnss', fns) = fold_rev (trans eqns') constrs + ((i, fname)::fnames, fnss, []) in - (fnameTs', (i, (fname, ls, rs, fns))::fnss') + (fnames', (i, (fname, ls, rs, fns))::fnss') end - | SOME fnameT' => - if fnameT = fnameT' then (fnameTs, fnss) + | SOME fname' => + if fname = fname' then (fnames, fnss) else raise RecError ("inconsistent functions for datatype " ^ quote tname)) end; @@ -195,18 +199,21 @@ (* make definition *) -fun make_def thy fs (fname, ls, rs, rec_name, tname) = +fun make_def ctxt fixes fs (fname, ls, rs, rec_name, tname) = let val used = map fst (fold Term.add_frees fs []); val x = (Name.variant used "x", dummyT); val frees = ls @ x :: rs; - val rhs = list_abs_free (frees, + val raw_rhs = list_abs_free (frees, list_comb (Const (rec_name, dummyT), fs @ [Free x])) - val def_name = Sign.base_name fname ^ "_" ^ Sign.base_name tname ^ "_def"; - val def_prop as _ $ _ $ t = - singleton (Syntax.check_terms (ProofContext.init thy)) - (Logic.mk_equals (Const (fname, dummyT), rhs)); - in ((def_name, def_prop), subst_bounds (rev (map Free frees), strip_abs_body t)) end; + val def_name = Thm.def_name (Sign.base_name fname); + val rhs = singleton (Syntax.check_terms ctxt) raw_rhs; + val SOME var = get_first (fn ((b, _), mx) => + if Binding.base_name b = fname then SOME (b, mx) else NONE) fixes; + in + ((var, ((Binding.name def_name, []), rhs)), + subst_bounds (rev (map Free frees), strip_abs_body rhs)) + end; (* find datatypes which contain all datatypes in tnames' *) @@ -227,27 +234,36 @@ local -fun gen_primrec_i note def alt_name invs fctxt eqns_atts thy = +fun prepare_spec prep_spec ctxt raw_fixes raw_spec = + let + val ((fixes, spec), _) = prep_spec + raw_fixes (map (single o apsnd single) raw_spec) ctxt + in (fixes, map (apsnd the_single) spec) end; + +fun gen_primrec set_group prep_spec prep_term invs fctxt raw_fixes raw_params raw_spec lthy = let - val (raw_eqns, atts) = split_list eqns_atts; - val eqns = map (apfst Binding.base_name) raw_eqns; - val dt_info = NominalPackage.get_nominal_datatypes thy; - val rec_eqns = fold_rev (process_eqn thy o snd) eqns []; + val (fixes', spec) = prepare_spec prep_spec lthy (raw_fixes @ raw_params) raw_spec; + val fixes = List.take (fixes', length raw_fixes); + val (names_atts, spec') = split_list spec; + val eqns' = map unquantify spec' + val eqns = fold_rev (process_eqn lthy (fn v => Variable.is_fixed lthy v + orelse exists (fn ((w, _), _) => v = Binding.base_name w) fixes)) spec' []; + val dt_info = NominalPackage.get_nominal_datatypes (ProofContext.theory_of lthy); val lsrs :: lsrss = maps (fn (_, (_, _, eqns)) => - map (fn (_, (ls, _, rs, _, _)) => ls @ rs) eqns) rec_eqns + map (fn (_, (ls, _, rs, _, _)) => ls @ rs) eqns) eqns val _ = (if forall (curry eq_set lsrs) lsrss andalso forall (fn (_, (_, _, (_, (ls, _, rs, _, _)) :: eqns)) => forall (fn (_, (ls', _, rs', _, _)) => ls = ls' andalso rs = rs') eqns - | _ => true) rec_eqns + | _ => true) eqns then () else primrec_err param_err); - val tnames = distinct (op =) (map (#1 o snd) rec_eqns); + val tnames = distinct (op =) (map (#1 o snd) eqns); val dts = find_dts dt_info tnames tnames; val main_fns = map (fn (tname, {index, ...}) => (index, - (fst o the o find_first (fn f => (#1 o snd) f = tname)) rec_eqns)) + (fst o the o find_first (fn (_, x) => #1 x = tname)) eqns)) dts; val {descr, rec_names, rec_rewrites, ...} = if null dts then @@ -256,32 +272,32 @@ val descr = map (fn (i, (tname, args, constrs)) => (i, (tname, args, map (fn (cname, cargs) => (cname, fold (fn (dTs, dT) => fn dTs' => dTs' @ dTs @ [dT]) cargs [])) constrs))) descr; - val (fnameTs, fnss) = - fold_rev (process_fun thy descr rec_eqns) main_fns ([], []); + val (fnames, fnss) = fold_rev (process_fun lthy descr eqns) main_fns ([], []); val (fs, defs) = fold_rev (get_fns fnss) (descr ~~ rec_names) ([], []); - val defs' = map (make_def thy fs) defs; - val nameTs1 = map snd fnameTs; - val nameTs2 = map fst rec_eqns; - val _ = if gen_eq_set (op =) (nameTs1, nameTs2) then () - else primrec_err ("functions " ^ commas_quote (map fst nameTs2) ^ - "\nare not mutually recursive"); - val primrec_name = - if alt_name = "" then (space_implode "_" (map (Sign.base_name o #1) defs)) else alt_name; - val (defs_thms', thy') = - thy - |> Sign.add_path primrec_name - |> fold_map def (map (fn ((name, t), _) => ((name, []), t)) defs'); - val cert = cterm_of thy'; + val defs' = map (make_def lthy fixes fs) defs; + val names1 = map snd fnames; + val names2 = map fst eqns; + val _ = if gen_eq_set (op =) (names1, names2) then () + else primrec_err ("functions " ^ commas_quote names2 ^ + "\nare not mutually recursive"); + val (defs_thms, lthy') = lthy |> + set_group ? LocalTheory.set_group (serial_string ()) |> + fold_map (apfst (snd o snd) oo + LocalTheory.define Thm.definitionK o fst) defs'; + val qualify = Binding.qualify + (space_implode "_" (map (Sign.base_name o #1) defs)); + val names_atts' = map (apfst qualify) names_atts; + val cert = cterm_of (ProofContext.theory_of lthy'); fun mk_idx eq = let - val Const c = head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop + val Free (name, _) = head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (Logic.strip_imp_concl eq)))); - val SOME i = AList.lookup op = (map swap fnameTs) c; + val SOME i = AList.lookup op = (map swap fnames) name; val SOME (_, _, constrs) = AList.lookup op = descr i; - val SOME (_, _, eqns) = AList.lookup op = rec_eqns c; + val SOME (_, _, eqns'') = AList.lookup op = eqns name; val SOME (cname, (_, cargs, _, _, _)) = find_first - (fn (_, (_, _, _, _, eq')) => eq = eq') eqns + (fn (_, (_, _, _, _, eq')) => eq = eq') eqns'' in (i, find_index (fn (cname', _) => cname = cname') constrs, cargs) end; val rec_rewritess = @@ -296,19 +312,15 @@ curry (List.take o swap) (length fvars) |> map cert; val invs' = (case invs of NONE => map (fn (i, _) => - let - val SOME (_, T) = AList.lookup op = fnameTs i - val (Ts, U) = strip_type T - in - Abs ("x", List.drop (Ts, length lsrs + 1) ---> U, HOLogic.true_const) - end) descr - | SOME invs' => invs'); + Abs ("x", fastype_of (snd (nth defs' i)), HOLogic.true_const)) descr + | SOME invs' => map (prep_term lthy') invs'); val inst = (map cert fvars ~~ cfs) @ (map (cert o Var) pvars ~~ map cert invs') @ (case ctxtvars of - [ctxtvar] => [(cert (Var ctxtvar), cert (the_default HOLogic.unit fctxt))] + [ctxtvar] => [(cert (Var ctxtvar), + cert (the_default HOLogic.unit (Option.map (prep_term lthy') fctxt)))] | _ => []); - val rec_rewrites' = map (fn (_, eq) => + val rec_rewrites' = map (fn eq => let val (i, j, cargs) = mk_idx eq val th = nth (nth rec_rewritess i) j; @@ -317,8 +329,8 @@ strip_comb |> snd in (cargs, Logic.strip_imp_prems eq, Drule.cterm_instantiate (inst @ - (map (cterm_of thy') cargs' ~~ map (cterm_of thy' o Free) cargs)) th) - end) eqns; + (map cert cargs' ~~ map (cert o Free) cargs)) th) + end) eqns'; val prems = foldr1 (common_prefix op aconv) (map (prems_of o #3) rec_rewrites'); val cprems = map cert prems; @@ -346,64 +358,37 @@ val rule = implies_intr_list rule_prems (Conjunction.intr_balanced (map mk_eqn (rec_rewrites' ~~ asmss))); - val goals = map (fn ((cargs, _, _), (_, eqn)) => - (list_all_free (cargs, eqn), [])) (rec_rewrites' ~~ eqns); + val goals = map (fn ((cargs, _, _), eqn) => + (list_all_free (cargs, eqn), [])) (rec_rewrites' ~~ eqns'); in - thy' |> - ProofContext.init |> + lthy' |> + Variable.add_fixes (map fst lsrs) |> snd |> Proof.theorem_i NONE - (fn thss => ProofContext.theory (fn thy => + (fn thss => fn goal_ctxt => let - val simps = map standard (flat thss); - val (simps', thy') = - fold_map note ((map fst eqns ~~ atts) ~~ map single simps) thy; - val simps'' = maps snd simps' + val simps = ProofContext.export goal_ctxt lthy' (flat thss); + val (simps', lthy'') = fold_map (LocalTheory.note Thm.theoremK) + (names_atts' ~~ map single simps) lthy' in - thy' - |> note (("simps", [Simplifier.simp_add]), simps'') + lthy'' + |> LocalTheory.note Thm.theoremK ((qualify (Binding.name "simps"), + [Attrib.internal (K Simplifier.simp_add)]), maps snd simps') |> snd - |> Sign.parent_path - end)) + end) [goals] |> Proof.apply (Method.Basic (fn _ => Method.RAW_METHOD (fn _ => - rewrite_goals_tac (map snd defs_thms') THEN + rewrite_goals_tac defs_thms THEN compose_tac (false, rule, length rule_prems) 1), Position.none)) |> Seq.hd end; -fun gen_primrec note def alt_name invs fctxt eqns thy = - let - val ((names, strings), srcss) = apfst split_list (split_list eqns); - val atts = map (map (Attrib.attribute thy)) srcss; - val eqn_ts = map (fn s => Syntax.read_prop_global thy s - handle ERROR msg => cat_error msg ("The error(s) above occurred for " ^ s)) strings; - val rec_ts = map (fn eq => head_of (fst (HOLogic.dest_eq - (HOLogic.dest_Trueprop (Logic.strip_imp_concl eq)))) - handle TERM _ => primrec_eq_err thy "not a proper equation" eq) eqn_ts; - val (_, eqn_ts') = OldPrimrecPackage.unify_consts thy rec_ts eqn_ts - in - gen_primrec_i note def alt_name - (Option.map (map (Syntax.read_term_global thy)) invs) - (Option.map (Syntax.read_term_global thy) fctxt) - (names ~~ eqn_ts' ~~ atts) thy - end; - -fun thy_note ((name, atts), thms) = - PureThy.add_thmss [((name, thms), atts)] #-> (fn [thms] => pair (name, thms)); -fun thy_def false ((name, atts), t) = - PureThy.add_defs false [((name, t), atts)] #-> (fn [thm] => pair (name, thm)) - | thy_def true ((name, atts), t) = - PureThy.add_defs_unchecked false [((name, t), atts)] #-> (fn [thm] => pair (name, thm)); - in -val add_primrec = gen_primrec thy_note (thy_def false); -val add_primrec_unchecked = gen_primrec thy_note (thy_def true); -val add_primrec_i = gen_primrec_i thy_note (thy_def false); -val add_primrec_unchecked_i = gen_primrec_i thy_note (thy_def true); +val add_primrec = gen_primrec false Specification.check_specification (K I); +val add_primrec_cmd = gen_primrec true Specification.read_specification Syntax.read_term; -end; (*local*) +end; (* outer syntax *) @@ -419,25 +404,26 @@ val parser2 = (invariant -- P.$$$ ":") |-- (Scan.repeat1 (unless_flag P.term) >> SOME) -- Scan.optional parser1 NONE || (parser1 >> pair NONE); -val parser3 = - unless_flag P.name -- Scan.optional parser2 (NONE, NONE) || - (parser2 >> pair ""); -val parser4 = - (P.$$$ "unchecked" >> K true) -- Scan.optional parser3 ("", (NONE, NONE)) || - (parser3 >> pair false); val options = Scan.optional (P.$$$ "(" |-- P.!!! - (parser4 --| P.$$$ ")")) (false, ("", (NONE, NONE))); + (parser2 --| P.$$$ ")")) (NONE, NONE); + +fun pipe_error t = P.!!! (Scan.fail_with (K + (cat_lines ["Equations must be separated by " ^ quote "|", quote t]))); -val primrec_decl = - options -- Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop); +val statement = SpecParse.opt_thm_name ":" -- P.prop --| Scan.ahead + ((P.term :-- pipe_error) || Scan.succeed ("","")); + +val statements = P.enum1 "|" statement; + +val primrec_decl = P.opt_target -- options -- + P.fixes -- P.for_fixes --| P.$$$ "where" -- statements; val _ = OuterSyntax.command "nominal_primrec" "define primitive recursive functions on nominal datatypes" K.thy_goal - (primrec_decl >> (fn ((unchecked, (alt_name, (invs, fctxt))), eqns) => - Toplevel.print o Toplevel.theory_to_proof - ((if unchecked then add_primrec_unchecked else add_primrec) alt_name invs fctxt - (map P.triple_swap eqns)))); + (primrec_decl >> (fn ((((opt_target, (invs, fctxt)), raw_fixes), raw_params), raw_spec) => + Toplevel.print o Toplevel.local_theory_to_proof opt_target + (add_primrec_cmd invs fctxt raw_fixes raw_params raw_spec))); end;