# HG changeset patch # User Christian Urban # Date 1308828505 -3600 # Node ID 0435c4dfd6f6014d9ba557aa5b3719188a78b025 # Parent eda5aeb056a64f91ddf5aaad67b17bdfaa5cd9be expanded the example diff -r eda5aeb056a6 -r 0435c4dfd6f6 Nominal/Ex/Classical.thy --- a/Nominal/Ex/Classical.thy Thu Jun 23 11:30:39 2011 +0100 +++ b/Nominal/Ex/Classical.thy Thu Jun 23 12:28:25 2011 +0100 @@ -8,13 +8,29 @@ atom_decl coname nominal_datatype trm = - Ax "name" "coname" -| Cut n::"coname" t1::"trm" c::"coname" t2::"trm" bind n in t1, bind c in t2 -| AndR c1::"coname" t1::"trm" c2::"coname" t2::"trm" "coname" bind c1 in t1, bind c2 in t2 -| AndL1 n::"name" t::"trm" "name" bind n in t -| AndL2 n::"name" t::"trm" "name" bind n in t -| ImpL c::"coname" t1::"trm" n::"name" t2::"trm" "name" bind c in t1, bind n in t2 -| ImpR c::"coname" n::"name" t::"trm" "coname" bind n c in t + Ax "name" "coname" +| Cut n::"coname" t1::"trm" c::"coname" t2::"trm" bind n in t1, bind c in t2 + ("Cut <_>._ '(_')._" [100,100,100,100] 100) +| NotR n::"name" t::"trm" "coname" bind n in t + ("NotR '(_')._ _" [100,100,100] 100) +| NotL c::"coname" t::"trm" "name" bind c in t + ("NotL <_>._ _" [100,100,100] 100) +| AndR c1::"coname" t1::"trm" c2::"coname" t2::"trm" "coname" bind c1 in t1, bind c2 in t2 + ("AndR <_>._ <_>._ _" [100,100,100,100,100] 100) +| AndL1 n::"name" t::"trm" "name" bind n in t + ("AndL1 '(_')._ _" [100,100,100] 100) +| AndL2 n::"name" t::"trm" "name" bind n in t + ("AndL2 '(_')._ _" [100,100,100] 100) +| OrR1 c::"coname" t::"trm" "coname" bind c in t + ("OrR1 <_>._ _" [100,100,100] 100) +| OrR2 c::"coname" t::"trm" "coname" bind c in t + ("OrR2 <_>._ _" [100,100,100] 100) +| OrL n1::"name" t1::"trm" n2::"name" t2::"trm" "name" bind n1 in t1, bind n2 in t2 + ("OrL '(_')._ '(_')._ _" [100,100,100,100,100] 100) +| ImpL c::"coname" t1::"trm" n::"name" t2::"trm" "name" bind c in t1, bind n in t2 + ("ImpL <_>._ '(_')._ _" [100,100,100,100,100] 100) +| ImpR c::"coname" n::"name" t::"trm" "coname" bind n c in t + ("ImpR '(_').<_>._ _" [100,100,100,100] 100) thm trm.distinct thm trm.induct @@ -30,8 +46,59 @@ thm trm.supp thm trm.supp[simplified] - - +nominal_primrec + crename :: "trm \ coname \ coname \ trm" ("_[_\c>_]" [100,100,100] 100) +where + "(Ax x a)[d\c>e] = (if a=d then Ax x e else Ax x a)" +| "atom a \ (d, e) \ (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)" +| "atom a \ (d, e) \ (NotL .M x)[d\c>e] = (NotL .(M[d\c>e]) x)" +| "\atom a \ (d, e); atom b \ (d, e)\ \ (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)" +| "(AndL1 (x).M y)[d\c>e] = AndL1 (x).(M[d\c>e]) y" +| "(AndL2 (x).M y)[d\c>e] = AndL2 (x).(M[d\c>e]) y" +| "atom a \ (d, e) \ (OrR1 .M b)[d\c>e] = + (if b=d then OrR1 .(M[d\c>e]) e else OrR1 .(M[d\c>e]) b)" +| "atom a \ (d, e) \ (OrR2 .M b)[d\c>e] = + (if b=d then OrR2 .(M[d\c>e]) e else OrR2 .(M[d\c>e]) b)" +| "(OrL (x).M (y).N z)[d\c>e] = OrL (x).(M[d\c>e]) (y).(N[d\c>e]) z" +| "atom a \ (d, e) \ (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)" +| "atom a \ (d, e) \ (ImpL .M (x).N y)[d\c>e] = ImpL .(M[d\c>e]) (x).(N[d\c>e]) y" + apply(simp only: eqvt_def) + apply(simp only: crename_graph_def) + apply (rule, perm_simp, rule) + apply(rule TrueI) + -- "covered all cases" + apply(case_tac x) + apply(rule_tac y="a" and c="(b, c)" in trm.strong_exhaust) + apply(simp) + apply(blast) + apply(simp add: fresh_star_def) + apply(blast) + apply(simp add: fresh_star_def) + apply(blast) + apply(simp add: fresh_star_def) + apply(blast) + apply(simp add: fresh_star_def) + apply(blast) + apply(simp add: fresh_star_def) + apply(blast) + apply(simp add: fresh_star_def) + apply(blast) + apply(simp add: fresh_star_def) + apply(blast) + apply(simp add: fresh_star_def) + apply(blast) + apply(simp add: fresh_star_def) + apply(metis) + apply(simp add: fresh_star_def) + apply(metis) + apply(simp add: fresh_star_def) + apply(metis) + -- "clashes" + apply(simp_all) + oops end