# HG changeset patch # User nipkow # Date 1660216980 -7200 # Node ID 3581dcee70dbc5166ca6f7e3d456a056076e57d2 # Parent dd04e81172a8b38cb41ea1dead1a31ff76f01dab removing the [simp] attribute breaks too many AFP entries severely diff -r dd04e81172a8 -r 3581dcee70db src/HOL/Library/NList.thy --- a/src/HOL/Library/NList.thy Thu Aug 11 11:57:19 2022 +0200 +++ b/src/HOL/Library/NList.thy Thu Aug 11 13:23:00 2022 +0200 @@ -14,20 +14,23 @@ lemma nlistsI: "\ size xs = n; set xs \ A \ \ xs \ nlists n A" by (simp add: nlists_def) -lemma nlistsE_length (*rm simp del*): "xs \ nlists n A \ size xs = n" +text \These [simp] attributes are double-edged. + Many proofs in Jinja rely on it but they can degrade performance.\ + +lemma nlistsE_length [simp]: "xs \ nlists n A \ size xs = n" by (simp add: nlists_def) lemma less_lengthI: "\ xs \ nlists n A; p < n \ \ p < size xs" -by (simp add: nlistsE_length) +by (simp) -lemma nlistsE_set[simp](*rm simp del*): "xs \ nlists n A \ set xs \ A" +lemma nlistsE_set[simp]: "xs \ nlists n A \ set xs \ A" unfolding nlists_def by (simp) lemma nlists_mono: assumes "A \ B" shows "nlists n A \ nlists n B" proof fix xs assume "xs \ nlists n A" - then obtain size: "size xs = n" and inA: "set xs \ A" by (simp add:nlistsE_length) + then obtain size: "size xs = n" and inA: "set xs \ A" by (simp) with assms have "set xs \ B" by simp with size show "xs \ nlists n B" by(clarsimp intro!: nlistsI) qed