diff -r 8ad8612e5d9b -r 09cf78bb53d4 Nominal/Ex/Lambda.thy --- a/Nominal/Ex/Lambda.thy Thu Jan 06 13:31:44 2011 +0000 +++ b/Nominal/Ex/Lambda.thy Thu Jan 06 14:02:10 2011 +0000 @@ -2,6 +2,7 @@ imports "../Nominal2" begin + atom_decl name nominal_datatype lam = @@ -19,161 +20,6 @@ thm lam.fv_bn_eqvt thm lam.size_eqvt -ML {* - Outer_Syntax.local_theory_to_proof; - Proof.theorem -*} - -ML {* -fun prove_strong_ind pred_name avoids ctxt = - let - val _ = () - in - Proof.theorem NONE (fn thss => fn ctxt => ctxt) [] ctxt - end - -(* outer syntax *) -local - structure P = Parse; - structure S = Scan - -in -val _ = - Outer_Syntax.local_theory_to_proof "nominal_inductive" - "prove strong induction theorem for inductive predicate involving nominal datatypes" - Keyword.thy_goal - (Parse.xname -- - (Scan.optional (Parse.$$$ "avoids" |-- Parse.enum1 "|" (Parse.name -- - (Parse.$$$ ":" |-- Parse.and_list1 Parse.term))) []) >> (fn (pred_name, avoids) => - prove_strong_ind pred_name avoids)) - -end -*} - - -section {* Typing *} - -nominal_datatype ty = - TVar string -| TFun ty ty ("_ \ _") - - -inductive - valid :: "(name \ ty) list \ bool" -where - "valid []" -| "\atom x \ Gamma; valid Gamma\ \ valid ((x, T)#Gamma)" - -inductive - typing :: "(name\ty) list \ lam \ ty \ bool" ("_ \ _ : _" [60,60,60] 60) -where - t_Var[intro]: "\valid \; (x, T) \ set \\ \ \ \ Var x : T" - | t_App[intro]: "\\ \ t1 : T1 \ T2; \ \ t2 : T1\ \ \ \ App t1 t2 : T2" - | t_Lam[intro]: "\atom x \ \; (x, T1) # \ \ t : T2\ \ \ \ Lam x t : T1 \ T2" - -equivariance valid -equivariance typing - -thm valid.eqvt -thm typing.eqvt -thm eqvts -thm eqvts_raw - -thm typing.induct[of "\" "t" "T", no_vars] - -(* -lemma - fixes c::"'a::fs" - assumes a: "\ \ t : T" - and a1: "\\ x T c. \valid \; (x, T) \ set \\ \ P c \ (Var x) T" - and a2: "\\ t1 T1 T2 t2 c. \\ \ t1 : T1 \ T2; \d. P d \ t1 T1 \ T2; \ \ t2 : T1; \d. P d \ t2 T1\ - \ P c \ (App t1 t2) T2" - and a3: "\x \ T1 t T2 c. \atom x \ c; atom x \ \; (x, T1) # \ \ t : T2; \d. P d ((x, T1) # \) t T2\ - \ P c \ (Lam x t) T1 \ T2" - shows "P c \ t T" -proof - - from a have "\p c. P c (p \ \) (p \ t) (p \ T)" - proof (induct) - case (t_Var \ x T p c) - then show ?case - apply - - apply(perm_strict_simp) - apply(rule a1) - apply(drule_tac p="p" in permute_boolI) - apply(perm_strict_simp add: permute_minus_cancel) - apply(assumption) - apply(rotate_tac 1) - apply(drule_tac p="p" in permute_boolI) - apply(perm_strict_simp add: permute_minus_cancel) - apply(assumption) - done - next - case (t_App \ t1 T1 T2 t2 p c) - then show ?case - apply - - apply(perm_strict_simp) - apply(rule a2) - apply(drule_tac p="p" in permute_boolI) - apply(perm_strict_simp add: permute_minus_cancel) - apply(assumption) - apply(assumption) - apply(rotate_tac 2) - apply(drule_tac p="p" in permute_boolI) - apply(perm_strict_simp add: permute_minus_cancel) - apply(assumption) - apply(assumption) - done - next - case (t_Lam x \ T1 t T2 p c) - then show ?case - apply - - apply(subgoal_tac "\q. (q \ {p \ atom x}) \* c \ - supp (p \ \, p \ Lam x t, p \ (T1 \ T2)) \* q") - apply(erule exE) - apply(rule_tac t="p \ \" and s="(q + p) \ \" in subst) - apply(simp only: permute_plus) - apply(rule supp_perm_eq) - apply(simp add: supp_Pair fresh_star_union) - apply(rule_tac t="p \ Lam x t" and s="(q + p) \ Lam x t" in subst) - apply(simp only: permute_plus) - apply(rule supp_perm_eq) - apply(simp add: supp_Pair fresh_star_union) - apply(rule_tac t="p \ (T1 \ T2)" and s="(q + p) \ (T1 \ T2)" in subst) - apply(simp only: permute_plus) - apply(rule supp_perm_eq) - apply(simp add: supp_Pair fresh_star_union) - apply(simp (no_asm) only: eqvts) - apply(rule a3) - apply(simp only: eqvts permute_plus) - apply(simp add: fresh_star_def) - apply(drule_tac p="q + p" in permute_boolI) - apply(perm_strict_simp add: permute_minus_cancel) - apply(assumption) - apply(rotate_tac 1) - apply(drule_tac p="q + p" in permute_boolI) - apply(perm_strict_simp add: permute_minus_cancel) - apply(assumption) - apply(drule_tac x="d" in meta_spec) - apply(drule_tac x="q + p" in meta_spec) - apply(perm_strict_simp add: permute_minus_cancel) - apply(assumption) - apply(rule at_set_avoiding2) - apply(simp add: finite_supp) - apply(simp add: finite_supp) - apply(simp add: finite_supp) - apply(rule_tac p="-p" in permute_boolE) - apply(perm_strict_simp add: permute_minus_cancel) - --"supplied by the user" - apply(simp add: fresh_star_prod) - apply(simp add: fresh_star_def) - sorry - qed - then have "P c (0 \ \) (0 \ t) (0 \ T)" . - then show "P c \ t T" by simp -qed - -*) - section {* Matching *}