diff -r 0c0b2031e42e -r 2cb49d09f059 src/HOL/HOLCF/LowerPD.thy --- a/src/HOL/HOLCF/LowerPD.thy Tue Dec 17 15:35:46 2024 +0100 +++ b/src/HOL/HOLCF/LowerPD.thy Tue Dec 17 23:07:13 2024 +0100 @@ -59,17 +59,33 @@ assumes 2: "\t u a. P (PDUnit a) t \ P (PDUnit a) (PDPlus t u)" assumes 3: "\t u v. \P t v; P u v\ \ P (PDPlus t u) v" shows "P t u" -using le -apply (induct t arbitrary: u rule: pd_basis_induct) -apply (erule rev_mp) -apply (induct_tac u rule: pd_basis_induct) -apply (simp add: 1) -apply (simp add: lower_le_PDUnit_PDPlus_iff) -apply (simp add: 2) -apply (subst PDPlus_commute) -apply (simp add: 2) -apply (simp add: lower_le_PDPlus_iff 3) -done + using le +proof (induct t arbitrary: u rule: pd_basis_induct) + case (PDUnit a) + then show ?case + proof (induct u rule: pd_basis_induct) + case PDUnit + then show ?case by (simp add: 1) + next + case (PDPlus t u) + from PDPlus(3) consider (t) "PDUnit a \\ t" | (u) "PDUnit a \\ u" + by (auto simp: lower_le_PDUnit_PDPlus_iff) + then show ?case + proof cases + case t + then have "P (PDUnit a) t" by (rule PDPlus(1)) + then show ?thesis by (rule 2) + next + case u + then have "P (PDUnit a) u" by (rule PDPlus(2)) + then have "P (PDUnit a) (PDPlus u t)" by (rule 2) + then show ?thesis by (simp only: PDPlus_commute) + qed + qed +next + case (PDPlus t t') + then show ?case by (simp add: lower_le_PDPlus_iff 3) +qed subsection \Type definition\ @@ -271,29 +287,42 @@ lemma lower_pd_induct1: assumes P: "adm P" assumes unit: "\x. P {x}\" - assumes insert: - "\x ys. \P {x}\; P ys\ \ P ({x}\ \\ ys)" + assumes insert: "\x ys. \P {x}\; P ys\ \ P ({x}\ \\ ys)" shows "P (xs::'a::bifinite lower_pd)" -apply (induct xs rule: lower_pd.principal_induct, rule P) -apply (induct_tac a rule: pd_basis_induct1) -apply (simp only: lower_unit_Rep_compact_basis [symmetric]) -apply (rule unit) -apply (simp only: lower_unit_Rep_compact_basis [symmetric] - lower_plus_principal [symmetric]) -apply (erule insert [OF unit]) -done +proof (induct xs rule: lower_pd.principal_induct) + have *: "P {Rep_compact_basis a}\" for a + by (rule unit) + show "P (lower_principal a)" for a + proof (induct a rule: pd_basis_induct1) + case PDUnit + from * show ?case + by (simp only: lower_unit_Rep_compact_basis [symmetric]) + next + case (PDPlus a t) + with * have "P ({Rep_compact_basis a}\ \\ lower_principal t)" + by (rule insert) + then show ?case + by (simp only: lower_unit_Rep_compact_basis [symmetric] lower_plus_principal [symmetric]) + qed +qed (rule P) -lemma lower_pd_induct - [case_names adm lower_unit lower_plus, induct type: lower_pd]: +lemma lower_pd_induct [case_names adm lower_unit lower_plus, induct type: lower_pd]: assumes P: "adm P" assumes unit: "\x. P {x}\" assumes plus: "\xs ys. \P xs; P ys\ \ P (xs \\ ys)" shows "P (xs::'a::bifinite lower_pd)" -apply (induct xs rule: lower_pd.principal_induct, rule P) -apply (induct_tac a rule: pd_basis_induct) -apply (simp only: lower_unit_Rep_compact_basis [symmetric] unit) -apply (simp only: lower_plus_principal [symmetric] plus) -done +proof (induct xs rule: lower_pd.principal_induct) + show "P (lower_principal a)" for a + proof (induct a rule: pd_basis_induct) + case PDUnit + then show ?case + by (simp only: lower_unit_Rep_compact_basis [symmetric] unit) + next + case PDPlus + then show ?case + by (simp only: lower_plus_principal [symmetric] plus) + qed +qed (rule P) subsection \Monadic bind\