diff -r bdec4a82f385 -r f76f187c91f9 src/HOL/Nominal/Examples/Crary.thy --- a/src/HOL/Nominal/Examples/Crary.thy Fri Feb 02 15:47:58 2007 +0100 +++ b/src/HOL/Nominal/Examples/Crary.thy Fri Feb 02 17:16:16 2007 +0100 @@ -102,7 +102,7 @@ v_nil[intro]: "valid []" | v_cons[intro]: "\valid \;a\\\ \ valid ((a,\)#\)" -lemma valid_eqvt: +lemma valid_eqvt[eqvt]: fixes pi:: "name prm" assumes a: "valid \" shows "valid (pi\\)" @@ -129,10 +129,10 @@ using assms by (induct) (auto) -lemma typing_eqvt : +lemma typing_eqvt[eqvt]: fixes pi::"name prm" assumes "\ \ t : T" - shows "(pi\\) \ (pi\t) : T" + shows "(pi\\) \ (pi\t) : (pi\T)" using assms apply(induct) apply(auto simp add: fresh_bij set_eqvt valid_eqvt) @@ -163,29 +163,28 @@ and x :: "'a::fs_name" assumes a: "\ \ e : T" and a1: "\\ x T c. \valid \; (x,T)\set \\ \ P c \ (Var x) T" - and a2: "\\ e1 T1 T2 e2 c. - \\ \ e1 : T1\T2 ; \c. P c \ e1 (T1\T2); \ \ e2 : T1 ; \c. P c \ e2 T1\ + and a2: "\\ e1 T1 T2 e2 c. \\c. P c \ e1 (T1\T2); \c. P c \ e2 T1\ \ P c \ (App e1 e2) T2" - and a3: "\x \ T1 t T2 c. - \x\(\,c); (x,T1)#\ \ t : T2 ; \c. P c ((x,T1)#\) t T2\ + and a3: "\x \ T1 t T2 c.\x\(\,c); \c. P c ((x,T1)#\) t T2\ \ P c \ (Lam [x].t) (T1\T2)" and a4: "\\ n c. valid \ \ P c \ (Const n) TBase" and a5: "\\ c. valid \ \ P c \ Unit TUnit" shows "P c \ e T" proof - - from a have "\(pi::name prm) c. P c (pi\\) (pi\e) T" + from a have "\(pi::name prm) c. P c (pi\\) (pi\e) (pi\T)" proof (induct) case (t_Var \ x T pi c) have "valid \" by fact - then have "valid (pi\\)" by (simp only: valid_eqvt) + then have "valid (pi\\)" by (simp only: eqvt) moreover have "(x,T)\set \" by fact then have "pi\(x,T)\pi\(set \)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst]) then have "(pi\x,T)\set (pi\\)" by (simp add: set_eqvt) - ultimately show "P c (pi\\) (pi\(Var x)) T" using a1 by simp + ultimately show "P c (pi\\) (pi\(Var x)) (pi\T)" using a1 by simp next case (t_App \ e1 T1 T2 e2 pi c) - thus "P c (pi\\) (pi\(App e1 e2)) T2" using a2 by (simp, blast intro: typing_eqvt) + thus "P c (pi\\) (pi\(App e1 e2)) (pi\T2)" using a2 + by (simp only: eqvt) (blast) next case (t_Lam x \ T1 t T2 pi c) obtain y::"name" where fs: "y\(pi\x,pi\\,pi\t,c)" by (erule exists_fresh[OF fs_name1]) @@ -194,11 +193,8 @@ have f0: "x\\" by fact have f1: "(pi\x)\(pi\\)" using f0 by (simp add: fresh_bij) have f2: "y\?pi'\\" by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp) - have pr1: "(x,T1)#\ \ t : T2" by fact - then have "(?pi'\((x,T1)#\)) \ (?pi'\t) : T2" by (simp only: typing_eqvt) - moreover - have ih1: "\c. P c (?pi'\((x,T1)#\)) (?pi'\t) T2" by fact - ultimately have "P c (?pi'\\) (Lam [y].(?pi'\t)) (T1\T2)" using fs f2 a3 + have ih1: "\c. P c (?pi'\((x,T1)#\)) (?pi'\t) (?pi'\T2)" by fact + then have "P c (?pi'\\) (Lam [y].(?pi'\t)) (T1\T2)" using fs f2 a3 by (simp add: calc_atm) then have "P c (?sw\pi\\) (?sw\(Lam [(pi\x)].(pi\t))) (T1\T2)" by (simp del: append_Cons add: calc_atm pt_name2) @@ -206,18 +202,17 @@ by (rule perm_fresh_fresh) (simp_all add: fs f1) moreover have "(?sw\(Lam [(pi\x)].(pi\t))) = Lam [(pi\x)].(pi\t)" by (rule perm_fresh_fresh) (simp_all add: fs f1 fresh_atm abs_fresh) - ultimately show "P c (pi\\) (pi\(Lam [x].t)) (T1\T2)" - by simp + ultimately show "P c (pi\\) (pi\(Lam [x].t)) (pi\T1\T2)" by simp next case (t_Const \ n pi c) - thus "P c (pi\\) (pi\(Const n)) (TBase)" using a4 by (simp, blast intro: valid_eqvt) + thus "P c (pi\\) (pi\(Const n)) (pi\TBase)" using a4 by (simp, blast intro: valid_eqvt) next case (t_Unit \ pi c) - thus "P c (pi\\) (pi\Unit) (TUnit)" using a5 by (simp, blast intro: valid_eqvt) + thus "P c (pi\\) (pi\Unit) (pi\TUnit)" using a5 by (simp, blast intro: valid_eqvt) qed - then have "P c (([]::name prm)\\) (([]::name prm)\e) T" by blast + then have "P c (([]::name prm)\\) (([]::name prm)\e) (([]::name prm)\T)" by blast then show "P c \ e T" by simp -qed +qed text {* capture-avoiding substitution *} @@ -279,7 +274,7 @@ and "Unit [y::=t'] = Unit" by (simp_all add: fresh_list_cons fresh_list_nil) -lemma subst_eqvt: +lemma subst_eqvt[eqvt]: fixes pi::"name prm" and t::"trm" shows "pi\(t[x::=t']) = (pi\t)[(pi\x)::=(pi\t')]" @@ -371,13 +366,13 @@ shows "valid \" using assms by (induct,auto elim:typing_valid) -lemma equiv_def_eqvt: +lemma equiv_def_eqvt[eqvt]: fixes pi::"name prm" assumes a: "\ \ s == t : T" - shows "(pi\\) \ (pi\s) == (pi\t) : T" + shows "(pi\\) \ (pi\s) == (pi\t) : (pi\T)" using a apply(induct) -apply(auto intro: typing_eqvt valid_eqvt simp add: fresh_bij subst_eqvt) +apply(auto intro: typing_eqvt valid_eqvt simp add: fresh_bij subst_eqvt simp del: perm_ty) apply(rule_tac x="pi\x" in Q_Ext) apply(simp add: fresh_bij)+ done @@ -387,46 +382,39 @@ fixes c::"'a::fs_name" assumes a0: "\ \ s == t : T" and a1: "\\ t T c. \ \ t : T \ P c \ t t T" - and a2: "\\ t s T c. \\ \ t == s : T; \d. P d \ t s T\ \ P c \ s t T" - and a3: "\\ s t T u c. - \\ \ s == t : T; \d. P d \ s t T; \ \ t == u : T; \d. P d \ t u T\ + and a2: "\\ t s T c. \\d. P d \ t s T\ \ P c \ s t T" + and a3: "\\ s t T u c. \\d. P d \ s t T; \d. P d \ t u T\ \ P c \ s u T" - and a4: "\x \ T1 s2 t2 T2 c. \x\\; x\c; (x,T1)#\ \ s2 == t2 : T2; \d. P d ((x,T1)#\) s2 t2 T2\ + and a4: "\x \ T1 s2 t2 T2 c. \x\\; x\c; \d. P d ((x,T1)#\) s2 t2 T2\ \ P c \ (Lam [x].s2) (Lam [x].t2) (T1\T2)" - and a5: "\\ s1 t1 T1 T2 s2 t2 c. - \\ \ s1 == t1 : T1\T2; \d. P d \ s1 t1 (T1\T2); - \ \ s2 == t2 : T1; \d. P d \ s2 t2 T1\ \ P c \ (App s1 s2) (App t1 t2) T2" + and a5: "\\ s1 t1 T1 T2 s2 t2 c. \\d. P d \ s1 t1 (T1\T2); \d. P d \ s2 t2 T1\ + \ P c \ (App s1 s2) (App t1 t2) T2" and a6: "\x \ T1 s12 t12 T2 s2 t2 c. - \x\(\,s2,t2); x\c; (x,T1)#\ \ s12 == t12 : T2; \d. P d ((x,T1)#\) s12 t12 T2; - \ \ s2 == t2 : T1; \d. P d \ s2 t2 T1\ + \x\(\,s2,t2); x\c; \d. P d ((x,T1)#\) s12 t12 T2; \d. P d \ s2 t2 T1\ \ P c \ (App (Lam [x].s12) s2) (t12[x::=t2]) T2" and a7: "\x \ s t T1 T2 c. - \x\(\,s,t); (x,T1)#\ \ App s (Var x) == App t (Var x) : T2; - \d. P d ((x,T1)#\) (App s (Var x)) (App t (Var x)) T2\ + \x\(\,s,t); \d. P d ((x,T1)#\) (App s (Var x)) (App t (Var x)) T2\ \ P c \ s t (T1\T2)" shows "P c \ s t T" proof - - from a0 have "\(pi::name prm) c. P c (pi\\) (pi\s) (pi\t) T" + from a0 have "\(pi::name prm) c. P c (pi\\) (pi\s) (pi\t) (pi\T)" proof(induct) case (Q_Refl \ t T pi c) - then show "P c (pi\\) (pi\t) (pi\t) T" using a1 by (simp add: typing_eqvt) + then show "P c (pi\\) (pi\t) (pi\t) (pi\T)" using a1 by (simp only: eqvt) next case (Q_Symm \ t s T pi c) - then show "P c (pi\\) (pi\s) (pi\t) T" using a2 by (simp only: equiv_def_eqvt) + then show "P c (pi\\) (pi\s) (pi\t) (pi\T)" using a2 by (simp only: equiv_def_eqvt) next case (Q_Trans \ s t T u pi c) - then show " P c (pi\\) (pi\s) (pi\u) T" using a3 by (blast intro: equiv_def_eqvt) + then show " P c (pi\\) (pi\s) (pi\u) (pi\T)" using a3 by (blast intro: equiv_def_eqvt) next case (Q_App \ s1 t1 T1 T2 s2 t2 pi c) - then show "P c (pi\\) (pi\App s1 s2) (pi\App t1 t2) T2" using a5 - by (simp, blast intro: equiv_def_eqvt) + then show "P c (pi\\) (pi\App s1 s2) (pi\App t1 t2) (pi\T2)" using a5 + by (simp only: eqvt) (blast) next case (Q_Ext x \ s t T1 T2 pi c) - then show "P c (pi\\) (pi\s) (pi\t) (T1\T2)" - apply(auto intro!: a7 simp add: fresh_bij) - apply(drule equiv_def_eqvt) - apply(simp) - done + then show "P c (pi\\) (pi\s) (pi\t) (pi\T1\T2)" + by (auto intro!: a7 simp add: fresh_bij) next case (Q_Abs x \ T1 s2 t2 T2 pi c) obtain y::"name" where fs: "y\(pi\x,pi\s2,pi\t2,pi\\,c)" by (rule exists_fresh[OF fs_name1]) @@ -435,13 +423,9 @@ have f1: "x\\" by fact have f2: "(pi\x)\(pi\\)" using f1 by (simp add: fresh_bij) have f3: "y\?pi'\\" using f1 by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp) - have pr1: "(x,T1)#\ \ s2 == t2 : T2" by fact - then have "?pi'\((x,T1)#\) \ (?pi'\s2) == (?pi'\t2) : T2" by (rule equiv_def_eqvt) - then have "((y,T1)#(?pi'\\)) \ (?pi'\s2) == (?pi'\t2) : T2" by (simp add: calc_atm) - moreover - have ih1: "\c. P c (?pi'\((x,T1)#\)) (?pi'\s2) (?pi'\t2) T2" by fact + have ih1: "\c. P c (?pi'\((x,T1)#\)) (?pi'\s2) (?pi'\t2) (?pi'\T2)" by fact then have "\c. P c ((y,T1)#(?pi'\\)) (?pi'\s2) (?pi'\t2) T2" by (simp add: calc_atm) - ultimately have "P c (?pi'\\) (?pi'\Lam [x].s2) (?pi'\Lam [x].t2) (T1\T2)" using a4 f3 fs + then have "P c (?pi'\\) (?pi'\Lam [x].s2) (?pi'\Lam [x].t2) (T1\T2)" using a4 f3 fs by (simp add: calc_atm) then have "P c (?sw\pi\\) (?sw\Lam [(pi\x)].(pi\s2)) (?sw\Lam [(pi\x)].(pi\t2)) (T1\T2)" by (simp del: append_Cons add: calc_atm pt_name2) @@ -452,7 +436,7 @@ moreover have "(?sw\(Lam [(pi\x)].(pi\t2))) = Lam [(pi\x)].(pi\t2)" by (rule perm_fresh_fresh) (simp_all add: fs f2 abs_fresh) ultimately have "P c (pi\\) (Lam [(pi\x)].(pi\s2)) (Lam [(pi\x)].(pi\t2)) (T1\T2)" by simp - then show "P c (pi\\) (pi\Lam [x].s2) (pi\Lam [x].t2) (T1\T2)" by simp + then show "P c (pi\\) (pi\Lam [x].s2) (pi\Lam [x].t2) (pi\T1\T2)" by simp next case (Q_Beta x \ s2 t2 T1 s12 t12 T2 pi c) obtain y::"name" where fs: "y\(pi\x,pi\s12,pi\t12,pi\s2,pi\t2,pi\\,c)" @@ -463,19 +447,12 @@ have f2: "(pi\x)\(pi\(\,s2,t2))" using f1 by (simp add: fresh_bij) have f3: "y\(?pi'\(\,s2,t2))" using f1 by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp fresh_prod) - have pr1: "(x,T1)#\ \ s12 == t12 : T2" by fact - then have "?pi'\((x,T1)#\) \ (?pi'\s12) == (?pi'\t12) : T2" by (rule equiv_def_eqvt) - then have "((y,T1)#(?pi'\\)) \ (?pi'\s12) == (?pi'\t12) : T2" by (simp add: calc_atm) - moreover - have ih1: "\c. P c (?pi'\((x,T1)#\)) (?pi'\s12) (?pi'\t12) T2" by fact - then have "\c. P c ((y,T1)#(?pi'\\)) (?pi'\s12) (?pi'\t12) T2" by (simp add: calc_atm) + have ih1: "\c. P c (?pi'\((x,T1)#\)) (?pi'\s12) (?pi'\t12) (?pi'\T2)" by fact + then have "\c. P c ((y,T1)#(?pi'\\)) (?pi'\s12) (?pi'\t12) (?pi'\T2)" by (simp add: calc_atm) moreover - have pr2: "\ \ s2 == t2 : T1" by fact - then have "(?pi'\\) \ (?pi'\s2) == (?pi'\t2) : T1" by (rule equiv_def_eqvt) - moreover - have ih2: "\c. P c (?pi'\\) (?pi'\s2) (?pi'\t2) T1" by fact - ultimately have "P c (?pi'\\) (?pi'\(App (Lam [x].s12) s2)) (?pi'\(t12[x::=t2])) T2" - using a6 f3 fs by (simp add: subst_eqvt calc_atm) + have ih2: "\c. P c (?pi'\\) (?pi'\s2) (?pi'\t2) (?pi'\T1)" by fact + ultimately have "P c (?pi'\\) (?pi'\(App (Lam [x].s12) s2)) (?pi'\(t12[x::=t2])) (?pi'\T2)" + using a6 f3 fs by (force simp add: subst_eqvt calc_atm del: perm_ty) then have "P c (?sw\pi\\) (?sw\(App (Lam [(pi\x)].(pi\s12)) (pi\s2))) (?sw\((pi\t12)[(pi\x)::=(pi\t2)])) T2" by (simp del: append_Cons add: calc_atm pt_name2 subst_eqvt) @@ -488,9 +465,9 @@ (simp_all add: fs[simplified] f2[simplified] fresh_subst fresh_subst'') ultimately have "P c (pi\\) (App (Lam [(pi\x)].(pi\s12)) (pi\s2)) ((pi\t12)[(pi\x)::=(pi\t2)]) T2" by simp - then show "P c (pi\\) (pi\App (Lam [x].s12) s2) (pi\t12[x::=t2]) T2" by (simp add: subst_eqvt) + then show "P c (pi\\) (pi\App (Lam [x].s12) s2) (pi\t12[x::=t2]) (pi\T2)" by (simp add: subst_eqvt) qed - then have "P c (([]::name prm)\\) (([]::name prm)\s) (([]::name prm)\t) T" by blast + then have "P c (([]::name prm)\\) (([]::name prm)\s) (([]::name prm)\t) (([]::name prm)\T)" by blast then show "P c \ s t T" by simp qed @@ -538,7 +515,7 @@ using a by (induct) (auto simp add: subst_eqvt fresh_bij) -lemma whn_eqvt: +lemma whn_eqvt[eqvt]: fixes pi::"name prm" assumes a: "t \ t'" shows "(pi\t) \ (pi\t')" @@ -568,10 +545,10 @@ | QAP_App[intro]: "\\ \ p \ q : T1 \ T2; \ \ s <=> t : T1 \ \ \ \ App p s \ App q t : T2" | QAP_Const[intro]: "valid \ \ \ \ Const n \ Const n : TBase" -lemma alg_equiv_alg_path_equiv_eqvt: +lemma alg_equiv_alg_path_equiv_eqvt[eqvt]: fixes pi::"name prm" - shows "\ \ s <=> t : T \ (pi\\) \ (pi\s) <=> (pi\t) : T" - and "\ \ p \ q : T \ (pi\\) \ (pi\p) \ (pi\q) : T" + shows "\ \ s <=> t : T \ (pi\\) \ (pi\s) <=> (pi\t) : (pi\T)" + and "\ \ p \ q : T \ (pi\\) \ (pi\p) \ (pi\q) : (pi\T)" apply(induct rule: alg_equiv_alg_path_equiv.inducts) apply(auto intro: whn_eqvt simp add: fresh_bij valid_eqvt) apply(rule_tac x="pi\x" in QAT_Arrow) @@ -606,7 +583,7 @@ then show "\c (pi::name prm). P1 c (pi\\) (pi\s) (pi\t) TBase" apply(auto) apply(rule_tac p="pi\q" and q="pi\p" in a1) - apply(simp_all add: whn_eqvt alg_equiv_alg_path_equiv_eqvt) + apply(simp_all add: whn_eqvt alg_equiv_alg_path_equiv_eqvt[simplified]) done next case (QAT_Arrow x \ s t T1 T2) @@ -623,8 +600,8 @@ by (simp only: pt_name2 fresh_left, auto simp add: perm_pi_simp calc_atm) then have f3': "y\?pi'\\" "y\?pi'\s" "y\?pi'\t" by (auto simp add: fresh_prod) have pr1: "(x,T1)#\ \ App s (Var x) <=> App t (Var x) : T2" by fact - then have "(?pi'\((x,T1)#\)) \ (?pi'\(App s (Var x))) <=> (?pi'\(App t (Var x))) : T2" - by (simp only: alg_equiv_alg_path_equiv_eqvt) + then have "(?pi'\((x,T1)#\)) \ (?pi'\(App s (Var x))) <=> (?pi'\(App t (Var x))) : (?pi'\T2)" + by (rule alg_equiv_alg_path_equiv_eqvt) then have "(y,T1)#(?pi'\\) \ (App (?pi'\s) (Var y)) <=> (App (?pi'\t) (Var y)) : T2" by (simp add: calc_atm) moreover @@ -659,7 +636,7 @@ next case (QAP_App \ p q T1 T2 s t) then show "\c (pi::name prm). P2 c (pi\\) (pi\App p s) (pi\App q t) T2" - by (auto intro!: a5 simp add: alg_equiv_alg_path_equiv_eqvt) + by (auto intro!: a5 simp add: alg_equiv_alg_path_equiv_eqvt[simplified]) next case (QAP_Const \ n) then show "\c (pi::name prm). P2 c (pi\\) (pi\Const n) (pi\Const n) TBase" @@ -724,7 +701,6 @@ proof (nominal_induct \ s t T and \ s t T avoiding: \' rule: alg_equiv_alg_path_equiv_strong_inducts) case (QAT_Arrow x \ s t T1 T2 \') have fs:"x\\" "x\s" "x\t" by fact - have h1:"(x,T1)#\ \ App s (Var x) <=> App t (Var x) : T2" by fact have h2:"\\\'" by fact have ih:"\\'. \(x,T1)#\ \ \'\ \ \' \ App s (Var x) <=> App t (Var x) : T2" by fact have "x\\'" by fact @@ -984,7 +960,8 @@ have "(pi\((x,T1)#\)) \ pi \ (App s (Var x)) <=> pi \ (App t (Var x)) : T2" by fact then have "(pi\((x,T1)#\)) \ (App (pi\s) (Var (pi\x))) <=> (App (pi\t) (Var (pi\x))) : T2" by auto moreover have "pi\((x,T1)#\) = (pi\x,pi\T1)#(pi\\)" by auto - ultimately have "((pi\x,T1)#(pi\\)) \ (App (pi\s) (Var (pi\x))) <=> (App (pi\t) (Var (pi\x))) : T2" using perm_ty by auto + ultimately have "((pi\x,T1)#(pi\\)) \ (App (pi\s) (Var (pi\x))) <=> (App (pi\t) (Var (pi\x))) : T2" + using perm_ty by auto then show ?case using h by auto qed (auto elim:whn_eqvt valid_eqvt) @@ -1271,7 +1248,6 @@ case (Q_Abs x \ T1 s2 t2 T2 \' \ \) have fs:"x\\" by fact have fs2: "x\\" "x\\" by fact - have h1:"(x,T1)#\ \ s2 == t2 : T2" by fact have h2:"\' \ \ is \ over \" by fact have ih:"\\' \ \. \' \ \ is \ over (x,T1)#\ \ \' \ \ is \ : T2" by fact {