diff -r 5abac261b5ce -r 0b692f37a771 Nominal/Ex/Lambda.thy --- a/Nominal/Ex/Lambda.thy Sun Apr 25 09:26:36 2010 +0200 +++ b/Nominal/Ex/Lambda.thy Mon Apr 26 08:08:20 2010 +0200 @@ -137,6 +137,98 @@ 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(rule_tac p="-p" in permute_boolE) + apply(perm_strict_simp add: permute_minus_cancel) + apply(assumption) + apply(rule_tac p="-p" in permute_boolE) + 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_tac ?T1.0="p \ T1" in a2) + apply(rule_tac p="-p" in permute_boolE) + apply(perm_strict_simp add: permute_minus_cancel) + apply(assumption) + apply(assumption) + apply(rule_tac p="-p" in permute_boolE) + 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(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(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(rule supp_perm_eq) + apply(simp add: supp_Pair fresh_star_union) + apply(simp add: eqvts) + apply(rule a3) + apply(simp add: fresh_star_def) + apply(rule_tac p="-q" in permute_boolE) + apply(perm_strict_simp add: permute_minus_cancel) + apply(rule_tac p="-p" in permute_boolE) + apply(perm_strict_simp add: permute_minus_cancel) + apply(assumption) + apply(rule_tac p="-q" in permute_boolE) + apply(perm_strict_simp add: permute_minus_cancel) + apply(rule_tac p="-p" in permute_boolE) + 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(simp) + 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) + + + done +(* HERE *) + + + + + + inductive tt :: "('a \ 'a \ bool) \ ('a \ 'a \ bool)" @@ -323,6 +415,26 @@ *) +ML {* + +fun prove_strong_ind (pred_name, avoids) ctxt = + Proof.theorem NONE (K I) [] ctxt + +local structure P = OuterParse and K = OuterKeyword in + +val _ = + OuterSyntax.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 + + end