diff -r 54aade5d0fe6 -r a9a1ed3f5023 Nominal/Ex/Lambda.thy --- a/Nominal/Ex/Lambda.thy Mon Jan 17 12:34:11 2011 +0000 +++ b/Nominal/Ex/Lambda.thy Mon Jan 17 12:37:37 2011 +0000 @@ -2,11 +2,8 @@ imports "../Nominal2" begin - atom_decl name -ML {* suffix *} - nominal_datatype lam = Var "name" | App "lam" "lam" @@ -22,460 +19,7 @@ thm lam.fv_bn_eqvt thm lam.size_eqvt -definition - "eqvt_at f x \ \p. p \ (f x) = f (p \ x)" -function - depth :: "lam \ nat" -where - "depth (Var x) = 1" -| "depth (App t1 t2) = (max (depth t1) (depth t2)) + 1" -| "depth (Lam x t) = (depth t) + 1" -oops - -section {* Matching *} - -definition - MATCH :: "('c::pt \ (bool * 'a::pt * 'b::pt)) \ 'b \ 'a \ 'b" -where - "MATCH M d x \ if (\!r. \q. M q = (True, x, r)) then (THE r. \q. M q = (True, x, r)) else d" - -(* -lemma MATCH_eqvt: - shows "p \ (MATCH M d x) = MATCH (p \ M) (p \ d) (p \ x)" -unfolding MATCH_def -apply(perm_simp the_eqvt) -apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *}) -apply(simp) -thm eqvts_raw -apply(subst if_eqvt) -apply(subst ex1_eqvt) -apply(subst permute_fun_def) -apply(subst ex_eqvt) -apply(subst permute_fun_def) -apply(subst eq_eqvt) -apply(subst permute_fun_app_eq[where f="M"]) -apply(simp only: permute_minus_cancel) -apply(subst permute_prod.simps) -apply(subst permute_prod.simps) -apply(simp only: permute_minus_cancel) -apply(simp only: permute_bool_def) -apply(simp) -apply(subst ex1_eqvt) -apply(subst permute_fun_def) -apply(subst ex_eqvt) -apply(subst permute_fun_def) -apply(subst eq_eqvt) - -apply(simp only: eqvts) -apply(simp) -apply(subgoal_tac "(p \ (\!r. \q. M q = (True, x, r))) = (\!r. \q. (p \ M) q = (True, p \ x, r))") -apply(drule sym) -apply(simp) -apply(rule impI) -apply(simp add: perm_bool) -apply(rule trans) -apply(rule pt_the_eqvt[OF pta at]) -apply(assumption) -apply(simp add: pt_ex_eqvt[OF pt at]) -apply(simp add: pt_eq_eqvt[OF ptb at]) -apply(rule cheat) -apply(rule trans) -apply(rule pt_ex1_eqvt) -apply(rule pta) -apply(rule at) -apply(simp add: pt_ex_eqvt[OF pt at]) -apply(simp add: pt_eq_eqvt[OF ptb at]) -apply(subst pt_pi_rev[OF pta at]) -apply(subst pt_fun_app_eq[OF pt at]) -apply(subst pt_pi_rev[OF pt at]) -apply(simp) -done - -lemma MATCH_cng: - assumes a: "M1 = M2" "d1 = d2" - shows "MATCH M1 d1 x = MATCH M2 d2 x" -using a by simp - -lemma MATCH_eq: - assumes a: "t = l x" "G x" "\x'. t = l x' \ G x' \ r x' = r x" - shows "MATCH (\x. (G x, l x, r x)) d t = r x" -using a -unfolding MATCH_def -apply(subst if_P) -apply(rule_tac a="r x" in ex1I) -apply(rule_tac x="x" in exI) -apply(blast) -apply(erule exE) -apply(drule_tac x="q" in meta_spec) -apply(auto)[1] -apply(rule the_equality) -apply(blast) -apply(erule exE) -apply(drule_tac x="q" in meta_spec) -apply(auto)[1] -done - -lemma MATCH_eq2: - assumes a: "t = l x1 x2" "G x1 x2" "\x1' x2'. t = l x1' x2' \ G x1' x2' \ r x1' x2' = r x1 x2" - shows "MATCH (\(x1,x2). (G x1 x2, l x1 x2, r x1 x2)) d t = r x1 x2" -sorry - -lemma MATCH_neq: - assumes a: "\x. t = l x \ G x \ False" - shows "MATCH (\x. (G x, l x, r x)) d t = d" -using a -unfolding MATCH_def -apply(subst if_not_P) -apply(blast) -apply(rule refl) -done - -lemma MATCH_neq2: - assumes a: "\x1 x2. t = l x1 x2 \ G x1 x2 \ False" - shows "MATCH (\(x1,x2). (G x1 x2, l x1 x2, r x1 x2)) d t = d" -using a -unfolding MATCH_def -apply(subst if_not_P) -apply(auto) -done -*) - -ML {* -fun mk_avoids ctxt params name set = - let - val (_, ctxt') = ProofContext.add_fixes - (map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params) ctxt; - fun mk s = - let - val t = Syntax.read_term ctxt' s; - val t' = list_abs_free (params, t) |> - funpow (length params) (fn Abs (_, _, t) => t) - in (t', HOLogic.dest_setT (fastype_of t)) end - handle TERM _ => - error ("Expression " ^ quote s ^ " to be avoided in case " ^ - quote name ^ " is not a set type"); - fun add_set p [] = [p] - | add_set (t, T) ((u, U) :: ps) = - if T = U then - let val S = HOLogic.mk_setT T - in (Const (@{const_name sup}, S --> S --> S) $ u $ t, T) :: ps - end - else (u, U) :: add_set (t, T) ps - in - (mk #> add_set) set - end; -*} - - -ML {* - writeln (commas (map (Syntax.string_of_term @{context} o fst) - (mk_avoids @{context} [] "t_Var" "{x}" []))) -*} - - -ML {* - -fun prove_strong_ind (pred_name, avoids) ctxt = - Proof.theorem NONE (K I) [] ctxt - -local structure P = Parse and K = Keyword in - -val _ = - Outer_Syntax.local_theory_to_proof "nominal_inductive" - "proves strong induction theorem for inductive predicate involving nominal datatypes" K.thy_goal - (P.xname -- (Scan.optional (P.$$$ "avoids" |-- P.enum1 "|" (P.name -- - (P.$$$ ":" |-- P.and_list1 P.term))) []) >> prove_strong_ind) - -end; - -*} - -(* -nominal_inductive typing -*) - -(* Substitution *) - -primrec match_Var_raw where - "match_Var_raw (Var_raw x) = Some x" -| "match_Var_raw (App_raw x y) = None" -| "match_Var_raw (Lam_raw n t) = None" - -quotient_definition - "match_Var :: lam \ name option" -is match_Var_raw - -lemma [quot_respect]: "(alpha_lam_raw ===> op =) match_Var_raw match_Var_raw" - apply rule - apply (induct_tac x y rule: alpha_lam_raw.induct) - apply simp_all - done - -lemmas match_Var_simps = match_Var_raw.simps[quot_lifted] - -primrec match_App_raw where - "match_App_raw (Var_raw x) = None" -| "match_App_raw (App_raw x y) = Some (x, y)" -| "match_App_raw (Lam_raw n t) = None" - -(* -quotient_definition - "match_App :: lam \ (lam \ lam) option" -is match_App_raw - -lemma [quot_respect]: - "(alpha_lam_raw ===> option_rel (prod_rel alpha_lam_raw alpha_lam_raw)) match_App_raw match_App_raw" - apply (intro fun_relI) - apply (induct_tac a b rule: alpha_lam_raw.induct) - apply simp_all - done - -lemmas match_App_simps = match_App_raw.simps[quot_lifted] - -definition new where - "new (s :: 'a :: fs) = (THE x. \a \ supp s. atom x \ a)" - -definition - "match_Lam (S :: 'a :: fs) t = (if (\n s. (t = Lam n s)) then - (let z = new (S, t) in Some (z, THE s. t = Lam z s)) else None)" - -lemma lam_half_inj: "(Lam z s = Lam z sa) = (s = sa)" - apply auto - apply (simp only: lam.eq_iff alphas) - apply clarify - apply (simp add: eqvts) - sorry - -lemma match_Lam_simps: - "match_Lam S (Var n) = None" - "match_Lam S (App l r) = None" - "z = new (S, (Lam z s)) \ match_Lam S (Lam z s) = Some (z, s)" - apply (simp_all add: match_Lam_def) - apply (simp add: lam_half_inj) - apply auto - done -*) -(* -lemma match_Lam_simps2: - "atom n \ ((S :: 'a :: fs), Lam n s) \ match_Lam S (Lam n s) = Some (n, s)" - apply (rule_tac t="Lam n s" - and s="Lam (new (S, (Lam n s))) ((n \ (new (S, (Lam n s)))) \ s)" in subst) - defer - apply (subst match_Lam_simps(3)) - defer - apply simp -*) - -(*primrec match_Lam_raw where - "match_Lam_raw (S :: atom set) (Var_raw x) = None" -| "match_Lam_raw S (App_raw x y) = None" -| "match_Lam_raw S (Lam_raw n t) = (let z = new (S \ (fv_lam_raw t - {atom n})) in Some (z, (n \ z) \ t))" - -quotient_definition - "match_Lam :: (atom set) \ lam \ (name \ lam) option" -is match_Lam_raw - -lemma swap_fresh: - assumes a: "fv_lam_raw t \* p" - shows "alpha_lam_raw (p \ t) t" - using a apply (induct t) - apply (simp add: supp_at_base fresh_star_def) - apply (rule alpha_lam_raw.intros) - apply (metis Rep_name_inverse atom_eqvt atom_name_def fresh_perm) - apply (simp) - apply (simp only: fresh_star_union) - apply clarify - apply (rule alpha_lam_raw.intros) - apply simp - apply simp - apply simp - apply (rule alpha_lam_raw.intros) - sorry - -lemma [quot_respect]: - "(op = ===> alpha_lam_raw ===> option_rel (prod_rel op = alpha_lam_raw)) match_Lam_raw match_Lam_raw" - proof (intro fun_relI, clarify) - fix S t s - assume a: "alpha_lam_raw t s" - show "option_rel (prod_rel op = alpha_lam_raw) (match_Lam_raw S t) (match_Lam_raw S s)" - using a proof (induct t s rule: alpha_lam_raw.induct) - case goal1 show ?case by simp - next - case goal2 show ?case by simp - next - case (goal3 x t y s) - then obtain p where "({atom x}, t) \gen (\x1 x2. alpha_lam_raw x1 x2 \ - option_rel (prod_rel op = alpha_lam_raw) (match_Lam_raw S x1) - (match_Lam_raw S x2)) fv_lam_raw p ({atom y}, s)" .. - then have - c: "fv_lam_raw t - {atom x} = fv_lam_raw s - {atom y}" and - d: "(fv_lam_raw t - {atom x}) \* p" and - e: "alpha_lam_raw (p \ t) s" and - f: "option_rel (prod_rel op = alpha_lam_raw) (match_Lam_raw S (p \ t)) (match_Lam_raw S s)" and - g: "p \ {atom x} = {atom y}" unfolding alphas(1) by - (elim conjE, assumption)+ - let ?z = "new (S \ (fv_lam_raw t - {atom x}))" - have h: "?z = new (S \ (fv_lam_raw s - {atom y}))" using c by simp - show ?case - unfolding match_Lam_raw.simps Let_def option_rel.simps prod_rel.simps split_conv - proof - show "?z = new (S \ (fv_lam_raw s - {atom y}))" by (fact h) - next - have "atom y \ p" sorry - have "fv_lam_raw t \* ((x \ y) \ p)" sorry - then have "alpha_lam_raw (((x \ y) \ p) \ t) t" using swap_fresh by auto - then have "alpha_lam_raw (p \ t) ((x \ y) \ t)" sorry - have "alpha_lam_raw t ((x \ y) \ s)" sorry - then have "alpha_lam_raw ((x \ ?z) \ t) ((y \ ?z) \ s)" using eqvts(15) sorry - then show "alpha_lam_raw ((x \ new (S \ (fv_lam_raw t - {atom x}))) \ t) - ((y \ new (S \ (fv_lam_raw s - {atom y}))) \ s)" unfolding h . - qed - qed - qed - -lemmas match_Lam_simps = match_Lam_raw.simps[quot_lifted] -*) -(* -lemma app_some: "match_App x = Some (a, b) \ x = App a b" -by (induct x rule: lam.induct) (simp_all add: match_App_simps) - -lemma lam_some: "match_Lam S x = Some (z, s) \ x = Lam z s \ atom z \ S" - apply (induct x rule: lam.induct) - apply (simp_all add: match_Lam_simps) - apply (thin_tac "match_Lam S lam = Some (z, s) \ lam = Lam z s \ atom z \ S") - apply (simp add: match_Lam_def) - apply (subgoal_tac "\n s. Lam name lam = Lam n s") - prefer 2 - apply auto[1] - apply (simp add: Let_def) - apply (thin_tac "\n s. Lam name lam = Lam n s") - apply clarify - apply (rule conjI) - apply (rule_tac t="THE s. Lam name lam = Lam (new (S, Lam name lam)) s" and - s="(name \ (new (S, Lam name lam))) \ lam" in subst) - defer - apply (simp add: lam.eq_iff) - apply (rule_tac x="(name \ (new (S, Lam name lam)))" in exI) - apply (simp add: alphas) - apply (simp add: eqvts) - apply (rule conjI) - sorry - -function subst where -"subst v s t = ( - case match_Var t of Some n \ if n = v then s else Var n | None \ - case match_App t of Some (l, r) \ App (subst v s l) (subst v s r) | None \ - case match_Lam (v,s) t of Some (n, t) \ Lam n (subst v s t) | None \ undefined)" -by pat_completeness auto - -termination apply (relation "measure (\(_, _, t). size t)") - apply auto[1] - apply (case_tac a) apply simp - apply (frule lam_some) apply simp - apply (case_tac a) apply simp - apply (frule app_some) apply simp - apply (case_tac a) apply simp - apply (frule app_some) apply simp -done - -lemmas lam_exhaust = lam_raw.exhaust[quot_lifted] - -lemma subst_eqvt: - "p \ (subst v s t) = subst (p \ v) (p \ s) (p \ t)" - proof (induct v s t rule: subst.induct) - case (1 v s t) - show ?case proof (cases t rule: lam_exhaust) - fix n - assume "t = Var n" - then show ?thesis by (simp add: match_Var_simps) - next - fix l r - assume "t = App l r" - then show ?thesis - apply (simp only:) - apply (subst subst.simps) - apply (subst match_Var_simps) - apply (simp only: option.cases) - apply (subst match_App_simps) - apply (simp only: option.cases) - apply (simp only: prod.cases) - apply (simp only: lam.perm) - apply (subst (3) subst.simps) - apply (subst match_Var_simps) - apply (simp only: option.cases) - apply (subst match_App_simps) - apply (simp only: option.cases) - apply (simp only: prod.cases) - apply (subst 1(2)[of "(l, r)" "l" "r"]) - apply (simp add: match_Var_simps) - apply (simp add: match_App_simps) - apply (rule refl) - apply (subst 1(3)[of "(l, r)" "l" "r"]) - apply (simp add: match_Var_simps) - apply (simp add: match_App_simps) - apply (rule refl) - apply (rule refl) - done - next - fix n t' - assume "t = Lam n t'" - then show ?thesis - apply (simp only: ) - apply (simp only: lam.perm) - apply (subst subst.simps) - apply (subst match_Var_simps) - apply (simp only: option.cases) - apply (subst match_App_simps) - apply (simp only: option.cases) - apply (rule_tac t="Lam n t'" and s="Lam (new ((v, s), Lam n t')) ((n \ new ((v, s), Lam n t')) \ t')" in subst) - defer - apply (subst match_Lam_simps) - defer - apply (simp only: option.cases) - apply (simp only: prod.cases) - apply (subst (2) subst.simps) - apply (subst match_Var_simps) - apply (simp only: option.cases) - apply (subst match_App_simps) - apply (simp only: option.cases) - apply (rule_tac t="Lam (p \ n) (p \ t')" and s="Lam (new ((p \ v, p \ s), Lam (p \ n) (p \ t'))) (((p \ n) \ new ((p \ v, p \ s), Lam (p \ n) (p \ t'))) \ t')" in subst) - defer - apply (subst match_Lam_simps) - defer - apply (simp only: option.cases) - apply (simp only: prod.cases) - apply (simp only: lam.perm) - thm 1(1) - sorry - qed - qed - -lemma subst_proper_eqs: - "subst y s (Var x) = (if x = y then s else (Var x))" - "subst y s (App l r) = App (subst y s l) (subst y s r)" - "atom x \ (t, s) \ subst y s (Lam x t) = Lam x (subst y s t)" - apply (subst subst.simps) - apply (simp only: match_Var_simps) - apply (simp only: option.simps) - apply (subst subst.simps) - apply (simp only: match_App_simps) - apply (simp only: option.simps) - apply (simp only: prod.simps) - apply (simp only: match_Var_simps) - apply (simp only: option.simps) - apply (subst subst.simps) - apply (simp only: match_Var_simps) - apply (simp only: option.simps) - apply (simp only: match_App_simps) - apply (simp only: option.simps) - apply (rule_tac t="Lam x t" and s="Lam (new ((y, s), Lam x t)) ((x \ new ((y, s), Lam x t)) \ t)" in subst) - defer - apply (subst match_Lam_simps) - defer - apply (simp only: option.simps) - apply (simp only: prod.simps) - sorry -*) end