diff -r dd874e6a3282 -r f274975039b2 src/HOL/Nominal/Examples/Crary.thy --- a/src/HOL/Nominal/Examples/Crary.thy Thu Jun 21 13:23:33 2007 +0200 +++ b/src/HOL/Nominal/Examples/Crary.thy Thu Jun 21 13:49:27 2007 +0200 @@ -319,12 +319,13 @@ declare trm.inject [simp add] declare ty.inject [simp add] -inductive_cases2 t_Lam_elim_auto[elim]: "\ \ Lam [x].t : T" -inductive_cases2 t_Var_elim_auto[elim]: "\ \ Var x : T" -inductive_cases2 t_App_elim_auto[elim]: "\ \ App x y : T" -inductive_cases2 t_Const_elim_auto[elim]: "\ \ Const n : T" -inductive_cases2 t_Unit_elim_auto[elim]: "\ \ Unit : TUnit" -inductive_cases2 t_Unit_elim_auto'[elim]: "\ \ s : TUnit" +inductive_cases2 t_inv_auto[elim]: + "\ \ Lam [x].t : T" + "\ \ Var x : T" + "\ \ App x y : T" + "\ \ Const n : T" + "\ \ Unit : TUnit" + "\ \ s : TUnit" declare trm.inject [simp del] declare ty.inject [simp del] @@ -366,15 +367,16 @@ declare trm.inject [simp add] declare ty.inject [simp add] -inductive_cases2 whr_Gen[elim]: "t \ t'" -inductive_cases2 whr_Lam[elim]: "Lam [x].t \ t'" -inductive_cases2 whr_App_Lam[elim]: "App (Lam [x].t12) t2 \ t" -inductive_cases2 whr_Var[elim]: "Var x \ t" -inductive_cases2 whr_Const[elim]: "Const n \ t" -inductive_cases2 whr_App[elim]: "App p q \ t" -inductive_cases2 whr_Const_Right[elim]: "t \ Const n" -inductive_cases2 whr_Var_Right[elim]: "t \ Var x" -inductive_cases2 whr_App_Right[elim]: "t \ App p q" +inductive_cases2 whr_inv_auto[elim]: + "t \ t'" + "Lam [x].t \ t'" + "App (Lam [x].t12) t2 \ t" + "Var x \ t" + "Const n \ t" + "App p q \ t" + "t \ Const n" + "t \ Var x" + "t \ App p q" declare trm.inject [simp del] declare ty.inject [simp del] @@ -421,12 +423,12 @@ shows "a=b" using a b apply (induct arbitrary: b) -apply (erule whr_App_Lam) +apply (erule whr_inv_auto(3)) apply (clarify) apply (rule subst_fun_eq) apply (simp) apply (force) -apply (erule whr_App) +apply (erule whr_inv_auto(6)) apply (blast)+ done @@ -468,23 +470,24 @@ declare trm.inject [simp add] declare ty.inject [simp add] -inductive_cases2 alg_equiv_Base_inv_auto[elim]: "\ \ s\t : TBase" -inductive_cases2 alg_equiv_Arrow_inv_auto[elim]: "\ \ s\t : T\<^isub>1 \ T\<^isub>2" - -inductive_cases2 alg_path_equiv_Base_inv_auto[elim]: "\ \ s\t : TBase" -inductive_cases2 alg_path_equiv_Unit_inv_auto[elim]: "\ \ s\t : TUnit" -inductive_cases2 alg_path_equiv_Arrow_inv_auto[elim]: "\ \ s\t : T\<^isub>1 \ T\<^isub>2" +inductive_cases2 alg_equiv_inv_auto[elim]: + "\ \ s\t : TBase" + "\ \ s\t : T\<^isub>1 \ T\<^isub>2" -inductive_cases2 alg_path_equiv_Var_left_inv_auto[elim]: "\ \ Var x \ t : T" -inductive_cases2 alg_path_equiv_Var_left_inv_auto'[elim]: "\ \ Var x \ t : T'" -inductive_cases2 alg_path_equiv_Var_right_inv_auto[elim]: "\ \ s \ Var x : T" -inductive_cases2 alg_path_equiv_Var_right_inv_auto'[elim]: "\ \ s \ Var x : T'" -inductive_cases2 alg_path_equiv_Const_left_inv_auto[elim]: "\ \ Const n \ t : T" -inductive_cases2 alg_path_equiv_Const_right_inv_auto[elim]: "\ \ s \ Const n : T" -inductive_cases2 alg_path_equiv_App_left_inv_auto[elim]: "\ \ App p s \ t : T" -inductive_cases2 alg_path_equiv_App_right_inv_auto[elim]: "\ \ s \ App q t : T" -inductive_cases2 alg_path_equiv_Lam_left_inv_auto[elim]: "\ \ Lam[x].s \ t : T" -inductive_cases2 alg_path_equiv_Lam_right_inv_auto[elim]: "\ \ t \ Lam[x].s : T" +inductive_cases2 alg_path_equiv_inv_auto[elim]: + "\ \ s\t : TBase" + "\ \ s\t : TUnit" + "\ \ s\t : T\<^isub>1 \ T\<^isub>2" + "\ \ Var x \ t : T" + "\ \ Var x \ t : T'" + "\ \ s \ Var x : T" + "\ \ s \ Var x : T'" + "\ \ Const n \ t : T" + "\ \ s \ Const n : T" + "\ \ App p s \ t : T" + "\ \ s \ App q t : T" + "\ \ Lam[x].s \ t : T" + "\ \ t \ Lam[x].s : T" declare trm.inject [simp del] declare ty.inject [simp del]