# HG changeset patch # User Christian Urban # Date 1396786044 -3600 # Node ID 7bc38b93a1fccac2d92ac6f9f28a2006d13a41ed # Parent 188826f1ccdb8ca22d9ee22f99e56b42162108b9 some small additions to examples diff -r 188826f1ccdb -r 7bc38b93a1fc Nominal/Ex/Lambda.thy --- a/Nominal/Ex/Lambda.thy Mon Mar 24 15:31:17 2014 +0000 +++ b/Nominal/Ex/Lambda.thy Sun Apr 06 13:07:24 2014 +0100 @@ -19,6 +19,39 @@ | App "lam" "lam" | Lam x::"name" l::"lam" binds x in l ("Lam [_]. _" [100, 100] 100) +lemma + "Lam [x]. Var x = Lam [y]. Var y" +proof - + obtain c::name where fresh: "atom c \ (Lam [x]. Var x, Lam [y]. Var y)" + by (metis obtain_fresh) + have "Lam [x]. Var x = (c \ x) \ Lam [x]. Var x" + using fresh by (rule_tac flip_fresh_fresh[symmetric]) (simp_all add: fresh_Pair) + also have "... = Lam [c].Var c" by simp + also have "... = (c \ y) \ Lam [c]. Var c" + using fresh by (rule_tac flip_fresh_fresh[symmetric]) (auto simp add: fresh_Pair) + also have "... = Lam [y]. Var y" by simp + finally show "Lam [x]. Var x = Lam [y]. Var y" . +qed + +definition + Name :: "nat \ name" +where + "Name n = Abs_name (Atom (Sort ''name'' []) n)" + +definition + "Ident2 = Lam [Name 1].(Var (Name 1))" + +definition + "Ident x = Lam [x].(Var x)" + +lemma "Ident2 = Ident x" +unfolding Ident_def Ident2_def +by simp + +lemma "Ident x = Ident y" +unfolding Ident_def +by simp + thm lam.strong_induct lemma alpha_lam_raw_eqvt[eqvt]: "p \ (alpha_lam_raw x y) = alpha_lam_raw (p \ x) (p \ y)" @@ -45,6 +78,7 @@ "is_app (Var x) = False" | "is_app (App t1 t2) = True" | "is_app (Lam [x]. t) = False" +thm is_app_graph_def is_app_graph_aux_def apply(simp add: eqvt_def is_app_graph_aux_def) apply(rule TrueI) apply(rule_tac y="x" in lam.exhaust) diff -r 188826f1ccdb -r 7bc38b93a1fc Nominal/Ex/TypeSchemes2.thy --- a/Nominal/Ex/TypeSchemes2.thy Mon Mar 24 15:31:17 2014 +0000 +++ b/Nominal/Ex/TypeSchemes2.thy Sun Apr 06 13:07:24 2014 +0100 @@ -68,6 +68,7 @@ "subst \ (Var X) = lookup \ X" | "subst \ (Fun T1 T2) = Fun (subst \ T1) (subst \ T2)" | "fset (map_fset atom xs) \* \ \ substs \ (All xs T) = All xs (subst \ T)" +thm subst_substs_graph_def subst_substs_graph_aux_def apply(simp add: subst_substs_graph_aux_def eqvt_def) apply(rule TrueI) apply (case_tac x)