diff -r a9f3600c9ae6 -r 589b1a0c75e6 Nominal/Ex/SFT/Utils.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/Ex/SFT/Utils.thy Fri Jun 24 10:30:06 2011 +0900 @@ -0,0 +1,101 @@ +header {* Utilities for defining constants and functions *} +theory Utils imports Lambda begin + +lemma beta_app: + "(\ x. M) \ V x \ M" + by (rule b3, rule bI) + (simp add: b1) + +lemma lam1_fast_app: + assumes leq: "\a. (L = \ a. (F (V a)))" + and su: "\x. atom x \ A \ F (V x) [x ::= A] = F A" + shows "L \ A \ F A" +proof - + obtain x :: var where a: "atom x \ A" using obtain_fresh by blast + show ?thesis + by (simp add: leq[of x], rule b3, rule bI, simp add: su b1 a) +qed + +lemma lam2_fast_app: + assumes leq: "\a b. a \ b \ L = \ a. \ b. (F (V a) (V b))" + and su: "\x y. atom x \ A \ atom y \ A \ atom x \ B \ atom y \ B \ + x \ y \ F (V x) (V y) [x ::= A] [y ::= B] = F A B" + shows "L \ A \ B \ F A B" +proof - + obtain x :: var where a: "atom x \ (A, B)" using obtain_fresh by blast + obtain y :: var where b: "atom y \ (x, A, B)" using obtain_fresh by blast + obtain z :: var where c: "atom z \ (x, y, A, B)" using obtain_fresh by blast + have *: "x \ y" "x \ z" "y \ z" + using a b c by (simp_all add: fresh_Pair fresh_at_base) blast+ + have ** : "atom y \ z" "atom x \ z" "atom y \ x" + "atom z \ y" "atom z \ x" "atom x \ y" + "atom x \ A" "atom y \ A" "atom z \ A" + "atom x \ B" "atom y \ B" "atom z \ B" + using a b c by (simp_all add: fresh_Pair fresh_at_base) blast+ + show ?thesis + apply (simp add: leq[OF *(1)]) + apply (rule b3) apply (rule b5) apply (rule bI) + apply (simp add: ** fresh_Pair) + apply (rule b3) apply (rule bI) apply (simp add: su b1 ** *) + done + qed + +lemma lam3_fast_app: + assumes leq: "\a b c. a \ b \ b \ c \ c \ a \ + L = \ a. \ b. \ c. (F (V a) (V b) (V c))" + and su: "\x y z. atom x \ A \ atom y \ A \ atom z \ A \ + atom x \ B \ atom y \ B \ atom z \ B \ + atom y \ B \ atom y \ B \ atom z \ B \ + x \ y \ y \ z \ z \ x \ + ((F (V x) (V y) (V z))[x ::= A] [y ::= B] [z ::= C] = F A B C)" + shows "L \ A \ B \ C \ F A B C" +proof - + obtain x :: var where a: "atom x \ (A, B, C)" using obtain_fresh by blast + obtain y :: var where b: "atom y \ (x, A, B, C)" using obtain_fresh by blast + obtain z :: var where c: "atom z \ (x, y, A, B, C)" using obtain_fresh by blast + have *: "x \ y" "y \ z" "z \ x" + using a b c by (simp_all add: fresh_Pair fresh_at_base) blast+ + have ** : "atom y \ z" "atom x \ z" "atom y \ x" + "atom z \ y" "atom z \ x" "atom x \ y" + "atom x \ A" "atom y \ A" "atom z \ A" + "atom x \ B" "atom y \ B" "atom z \ B" + "atom x \ C" "atom y \ C" "atom z \ C" + using a b c by (simp_all add: fresh_Pair fresh_at_base) blast+ + show ?thesis + apply (simp add: leq[OF *(1) *(2) *(3)]) + apply (rule b3) apply (rule b5) apply (rule b5) apply (rule bI) + apply (simp add: ** fresh_Pair) + apply (rule b3) apply (rule b5) apply (rule bI) + apply (simp add: ** fresh_Pair) + apply (rule b3) apply (rule bI) apply (simp add: su b1 ** *) + done + qed + +definition cn :: "nat \ var" where "cn n \ Abs_var (Atom (Sort ''Lambda.var'' []) n)" + +lemma cnd[simp]: "m \ n \ cn m \ cn n" + unfolding cn_def using Abs_var_inject by simp + +definition cx :: var where "cx \ cn 0" +definition cy :: var where "cy \ cn 1" +definition cz :: var where "cz \ cn 2" + +lemma cx_cy_cz[simp]: + "cx \ cy" "cx \ cz" "cz \ cy" + unfolding cx_def cy_def cz_def + by simp_all + +lemma noteq_fresh: "atom x \ y = (x \ y)" + by (simp add: fresh_at_base(2)) + +lemma fresh_fun_eqvt_app2: + assumes a: "eqvt f" + and b: "a \ x" "a \ y" + shows "a \ f x y" + using fresh_fun_eqvt_app[OF a b(1)] a b + by (metis fresh_fun_app) + +end + + +