diff -r 4436039cc5e1 -r 1b39ba5db2c1 Nominal/Ex/SFT/Consts.thy --- a/Nominal/Ex/SFT/Consts.thy Fri Jul 22 11:52:12 2011 +0100 +++ b/Nominal/Ex/SFT/Consts.thy Sun Jul 24 07:54:54 2011 +0200 @@ -94,13 +94,7 @@ by (rule, perm_simp, rule) qed -termination - by (relation "measure (\(t). size t)") - (simp_all add: lam.size) - -lemma numeral_eqvt[eqvt]: "p \ \x\ = \p \ x\" - by (induct x rule: lam.induct) - (simp_all add: Var_App_Abs_eqvt) +termination (eqvt) by lexicographic_order lemma supp_numeral[simp]: "supp \x\ = supp x" @@ -145,27 +139,7 @@ apply (simp_all add: fresh_at_base lam.fresh eqvt_def eqvts_raw fresh_rev) done -termination - by (relation "measure (\t. size t)") - (simp_all add: lam.size) - -lemma ltgt_eqvt[eqvt]: - "p \ \t\ = \p \ t\" -proof - - obtain x :: var where "atom x \ (t, p \ t)" using obtain_fresh by auto - then have *: "atom x \ t" "atom x \ (p \ t)" using fresh_Pair by simp_all - then show ?thesis using *[unfolded fresh_def] - apply (simp add: Abs1_eq_iff lam.fresh app_lst_eqvt Ltgt.simps) - apply (case_tac "p \ x = x") - apply (simp_all add: eqvts) - apply rule - apply (subst swap_fresh_fresh) - apply (simp_all add: fresh_at_base_permute_iff fresh_def[symmetric] fresh_at_base) - apply (subgoal_tac "eqvt app_lst") - apply (erule fresh_fun_eqvt_app2) - apply (simp_all add: fresh_at_base lam.fresh eqvt_def eqvts_raw fresh_rev) - done -qed +termination (eqvt) by lexicographic_order lemma ltgt_eq_iff[simp]: "\M\ = \N\ \ M = N" @@ -240,7 +214,7 @@ "F1 = \x. (App \ \Var\ \ (Var \ V x))" "a \ b \ a \ c \ c \ b \ F2 = \a. \b. \c. ((App \ (App \ \App\ \ (V c \ V a))) \ (V c \ V b))" "a \ b \ a \ x \ x \ b \ F3 = \a. \b. (App \ \Abs\ \ (Abs \ (\x. (V b \ (V a \ V x)))))" - apply (simp_all add: F1_def F2_def F3_def Abs1_eq_iff lam.fresh supp_at_base Var_App_Abs_eqvt numeral_eqvt flip_def[symmetric] fresh_at_base) + apply (simp_all add: F1_def F2_def F3_def Abs1_eq_iff lam.fresh supp_at_base Var_App_Abs_eqvt Numeral.eqvt flip_def[symmetric] fresh_at_base) apply (smt cx_cy_cz permute_flip_at)+ done @@ -291,7 +265,7 @@ "x \ y \ A1 = \x. \y. (F1 \ V x)" "a \ b \ a \ c \ c \ b \ A2 = \a. \b. \c. (F2 \ V a \ V b \ \[V c]\)" "a \ b \ A3 = \a. \b. (F3 \ V a \ \[V b]\)" - apply (simp_all add: Lam_A1_pre Lam_A2_pre Lam_A3_pre Abs1_eq_iff lam.fresh supp_at_base Var_App_Abs_eqvt numeral_eqvt flip_def[symmetric] fresh_at_base F_eqvt ltgt_eqvt) + apply (simp_all add: Lam_A1_pre Lam_A2_pre Lam_A3_pre Abs1_eq_iff lam.fresh supp_at_base Var_App_Abs_eqvt Numeral.eqvt flip_def[symmetric] fresh_at_base F_eqvt Ltgt.eqvt) apply (smt cx_cy_cz permute_flip_at)+ done