diff --git a/thys/PAPP_Impossibility/#Anonymous_PAPP_Lowering.thy# b/thys/PAPP_Impossibility/#Anonymous_PAPP_Lowering.thy# deleted file mode 100644 --- a/thys/PAPP_Impossibility/#Anonymous_PAPP_Lowering.thy# +++ /dev/null @@ -1,1309 +0,0 @@ -(* - File: Anonymous_PAPP_Lowering.thy - Author: Manuel Eberl, University of Innsbruck -*) -section \Lowering P-APP Rules to Smaller Settings\ -theory Anonymous_PAPP_Lowering - imports Anonymous_PAPP -begin - -text \ - In this section, we prove a number of lemmas (corresponding to Lemma~1 in the paper) that allow - us to take an anonymous P-APP rule with some additional properties (typically - Cardinality-Strategyproofness and Weak Representation or Weak Proportional Representation) and - construct from it an anonymous P-APP rule for a different setting, i.e.\ different number of - voters, parties, and/or result committee size. - - In the reverse direction, this also allows us to lift impossibility results from one setting - to another. -\ - -subsection \Preliminary Lemmas\ - -context card_stratproof_anon_papp -begin - -text \ - The following lemma is obtained by applying Strategyproofness repeatedly. It shows that - if we have \l\ voters with identical approval lists, then this entire group of voters has no - incentive to submit wrong preferences. That is, the outcome they obtain by submitting their - genuine approval lists is weakly preferred by them over all outcomes obtained where these \l\ - voters submit any other preferences (and the remaining $n - l$ voters submit the same preferences - as before). - - This is stronger than regular Strategyproofness, where we only demand that no voter has an - incentive to submit wrong preferences \<^emph>\unilaterally\ (and everyone else keeps the same - preferences). Here we know that the entire group of \l\ voters has no incentive to submit - wrong preferences in coordination with one another. -\ -lemma proposition2: - assumes "size B = l" "size A + l = n_voters" - assumes "X \ {}" "X \ parties" "{} \# A+B" "\X'\#A+B. X' \ parties" - shows "r (replicate_mset l X + A) \[Comm(X)] r (B + A)" - using assms -proof (induction l arbitrary: A B) - case 0 - thus ?case - by simp -next - case (Suc l A B) - from Suc.prems have "set_mset B \ {}" - by auto - then obtain Y where Y: "Y \# B" - by blast - define B' where "B' = B - {#Y#}" - define A' where "A' = A + {#Y#}" - have [simp]: "size B' = l" - using Suc.prems Y by (simp add: B'_def size_Diff_singleton) - have [simp]: "size A' = n_voters - l" - using Suc.prems Y by (simp add: A'_def) - - have "r (B' + A') \[Comm(X)] r (replicate_mset l X + A')" - by (rule Suc.IH) (use Suc.prems Y in \auto simp: A'_def B'_def size_Diff_singleton\) - also have "B' + A' = B + A" - using Y by (simp add: B'_def A'_def) - also have "r (replicate_mset l X + A') \[Comm(X)] r (replicate_mset (Suc l) X + A)" - proof (rule not_manipulable'') - show "replicate_mset (Suc l) X + A + {#Y#} = replicate_mset l X + A' + {#X#}" - by (simp add: A'_def) - next - show "is_pref_profile (replicate_mset (Suc l) X + A)" - using Suc.prems by unfold_locales (auto split: if_splits) - next - show "is_pref_profile (replicate_mset l X + A')" - using Suc.prems Y by unfold_locales (auto split: if_splits simp: A'_def) - qed - finally show ?case . -qed - -end - - -context card_stratproof_weak_rep_anon_papp -begin - -text \ - In a setting with Weak Representation and Cardinality-Strategyproofness, Proposition 2 allows - us to strengthen Weak Representation in the following way: Suppose we at least - $l \lfloor n/k\rfloor$ voters with the same approval list \X\, and \X\ consists of at least \l\ - parties. Then at least \l\ of the members of the result committee are in \X\. -\ -lemma proposition3: - assumes "is_pref_profile A" "X \ parties" "card X \ l" - assumes "committee_size > 0" - assumes "count A X \ l * \n_voters / committee_size\" - shows "size {# x\# r A. x \ X #} \ l" - using assms -proof (induction l arbitrary: A X rule: less_induct) - case (less l A X) - interpret A: anon_papp_profile n_voters parties committee_size A - by fact - consider "l = 0" | "l = 1" | "l > 1" - by force - thus ?case - proof cases - assume "l = 0" - thus ?thesis by simp - next - assume [simp]: "l = 1" - define n where "n = count A X" - with less.prems have "X \ {}" - by auto - then obtain x where x: "x \ X" - by blast - have "n \ size A" - unfolding n_def by (rule count_le_size) - hence "n \ n_voters" - by (simp add: A.size_A) - - have "count A X > 0" - by (rule Nat.gr0I) (use n_voters_pos less.prems in \auto simp: field_simps\) - hence "X \# A" - by force - have [simp]: "replicate_mset n X \# A" - by (simp add: n_def flip: count_le_replicate_mset_subset_eq) - - define A'' where "A'' = A - replicate_mset n X" - define A' where "A' = A'' + replicate_mset n {x}" - interpret A': anon_papp_profile n_voters parties committee_size A' - using A.A_nonempty A.A_subset A.size_A x \X \# A\ - by unfold_locales - (fastforce simp: A'_def A''_def size_Diff_submset subset_mset.add_increasing2 - split: if_splits dest!: in_diffD)+ - - have "x \# r A'" - proof (rule weak_representation) - show "is_pref_profile A'" - by (fact A'.anon_papp_profile_axioms) - next - have "n_voters \ committee_size * n" - using less.prems by (simp add: n_def ceiling_le_iff field_simps flip: of_nat_mult) - also have "n \ count A' {x}" - by (simp add: A'_def) - finally show "n_voters \ committee_size * count A' {x}" - by simp - qed - hence "1 \ count (r A') x" - by simp - also have "\ = size {# y \# r A'. y = x #}" - by simp - also have "\ \ size {# y \# r A'. y \ X #}" - by (intro size_mset_mono multiset_filter_mono') (use x in auto) - - also have "r A' \[Comm(X)] r A" - proof - - have "r (replicate_mset n {x} + A'') \[Comm(X)] r (replicate_mset n X + A'')" - proof (rule proposition2) - show "{} \# A'' + replicate_mset n {x}" - using A'.A_nonempty by (auto simp: A'_def) - show "\X'\#A'' + replicate_mset n {x}. X' \ parties" - using A'.A_subset x by (auto simp: A'_def dest: in_diffD) - show "size A'' + n = n_voters" - using \n \ n_voters\ by (auto simp: A''_def size_Diff_submset A.size_A) - qed (use less.prems in auto) - also have "replicate_mset n X + A'' = A" - by (simp add: A''_def n_def flip: count_le_replicate_mset_subset_eq) - finally show ?thesis - by (simp add: A'_def add_ac) - qed - hence "size {# y \# r A'. y \ X #} \ size {# y \# r A. y \ X #}" - by (simp add: committee_preference_def) - - finally show ?thesis - by simp - - next - assume l: "l > 1" - - define n where "n = count A X" - have "n \ size A" - unfolding n_def by (rule count_le_size) - hence "n \ n_voters" - by (simp add: A.size_A) - - define m where "m = nat (ceiling (n_voters / committee_size))" - have "n_voters / committee_size \ m" - unfolding m_def by linarith - hence m: "n_voters \ committee_size * m" - using \committee_size > 0\ by (simp add: field_simps flip: of_nat_mult) - have "real n_voters / real committee_size > 0" - using n_voters_pos less.prems by auto - hence m': "\real n_voters / real committee_size\ = int m" - by (simp add: m_def) - have "1 * m \ l * m" - using l by (intro mult_right_mono) auto - also have "l * m \ n" - using less.prems by (simp add: m' n_def flip: of_nat_mult) - finally have "m \ n" - by simp - - with less.prems l have "X \ {}" - by auto - then obtain x where x: "x \ X" - by blast - have "card (X - {x}) > 0" - using less.prems x l by simp - hence "X - {x} \ {}" - by force - - have "count A X > 0" - by (rule Nat.gr0I) (use n_voters_pos less.prems l in \auto simp: field_simps mult_le_0_iff\) - hence "X \# A" - by force - have [simp]: "replicate_mset n X \# A" - by (simp add: n_def flip: count_le_replicate_mset_subset_eq) - - define A'' where "A'' = A - replicate_mset n X" - define A' where "A' = A'' + replicate_mset m {x} + replicate_mset (n - m) (X - {x})" - interpret A': anon_papp_profile n_voters parties committee_size A' - proof - show "Y \ parties" if "Y \# A'" for Y - using that A.A_subset x \X \# A\ - by (fastforce simp: A'_def A''_def dest!: in_diffD split: if_splits) - next - show "{} \# A'" - using A.A_nonempty x \X \# A\ \X - {x} \ {}\ - by (auto simp: A'_def A''_def dest!: in_diffD split: if_splits) - next - show "size A' = n_voters" - using \m \ n\ - by (auto simp: A'_def A''_def A.size_A subset_mset.add_increasing2 size_Diff_submset) - qed - - have "x \# r A'" - proof (rule weak_representation) - show "is_pref_profile A'" - by (fact A'.anon_papp_profile_axioms) - next - have "n_voters \ committee_size * m" - by (fact m) - also have "m \ count A' {x}" - by (simp add: A'_def) - finally show "n_voters \ committee_size * count A' {x}" - by simp - qed - hence "1 \ count (r A') x" - by simp - also have "\ = size {# y \# r A'. y = x #}" - by simp - finally have 1: "size {#y \# r A'. y = x#} \ 1" . - - have 2: "size {# y \# r A'. y \ X - {x} #} \ l - 1" - proof (rule less.IH) - have "int (l - 1) * \real n_voters / real committee_size\ = int ((l - 1) * m)" - by (auto simp add: m' not_less) - also have "(l - 1) * m = l * m - m" - by (simp add: algebra_simps) - also have "l * m \ n" - using less.prems by (simp add: m' n_def flip: of_nat_mult) - hence "l * m - m \ n - m" - by (meson diff_le_mono) - also have "n - m \ count A' (X - {x})" - by (simp add: A'_def A''_def) - finally show "int (l - 1) * \real n_voters / real committee_size\ \ int (count A' (X - {x}))" - by simp - qed (use l A'.anon_papp_profile_axioms x less.prems in \auto\) - - have "1 + (l - 1) \ size {#y \# r A'. y = x#} + size {#y \# r A'. y \ X - {x}#}" - by (intro add_mono 1 2) - also have "\ = size ({#y \# r A'. y = x#} + {#y \# r A'. y \ X - {x}#})" - by simp - also have "{#y \# r A'. y = x#} + {#y \# r A'. y \ X - {x}#} = - {#y \# r A'. y = x \ y \ X - {x}#}" - by (rule filter_mset_disjunction [symmetric]) auto - also have "(\y. y = x \ y \ X - {x}) = (\y. y \ X)" - using x by auto - also have "1 + (l - 1) = l" - using l by simp - - also have "r A' \[Comm(X)] r A" - proof - - have "r (replicate_mset m {x} + replicate_mset (n - m) (X - {x}) + A'') \[Comm(X)] - r (replicate_mset n X + A'')" - proof (rule proposition2) - show " {} \# A'' + (replicate_mset m {x} + replicate_mset (n - m) (X - {x}))" - using A'.A_nonempty by (auto simp: A'_def) - show "\X'\# A'' + (replicate_mset m {x} + replicate_mset (n - m) (X - {x})). X' \ parties" - using A'.A_subset x by (auto simp: A'_def dest: in_diffD) - show "size A'' + n = n_voters" - using \n \ n_voters\ by (auto simp: A''_def size_Diff_submset A.size_A) - qed (use less.prems l \m \ n\ in auto) - also have "replicate_mset n X + A'' = A" - by (simp add: A''_def n_def flip: count_le_replicate_mset_subset_eq) - finally show ?thesis - by (simp add: A'_def add_ac) - qed - hence "size {# y \# r A'. y \ X #} \ size {# y \# r A. y \ X #}" - by (simp add: committee_preference_def) - - finally show ?thesis - by simp - qed -qed - -end - - -subsection \Dividing the number of voters\ - -text \ - If we have a PAPP rule that satisfies weak representation and cardinality strategyproofness, - for $ln$ voters, we can turn it into one for \n\ voters. - This is done by simply cloning each voter \l\ times. - - Consequently, if we have an impossibility result for $n$ voters, it also holds for any - integer multiple of \n\. -\ -locale divide_voters_card_stratproof_weak_rep_anon_papp = - card_stratproof_weak_rep_anon_papp "l * n_voters" parties committee_size r - for l n_voters parties committee_size r -begin - -definition lift_profile :: "'a set multiset \ 'a set multiset" where - "lift_profile A = (\X\#A. replicate_mset l X)" - -sublocale lowered: anon_papp_election n_voters parties - by standard (use n_voters_pos in auto) - -lemma l_pos: "l > 0" - using n_voters_pos by auto - -lemma empty_in_lift_profile_iff [simp]: "{} \# lift_profile A \ {} \# A" - using l_pos by (auto simp: lift_profile_def) - -lemma set_mset_lift_profile [simp]: "set_mset (lift_profile A) = set_mset A" - using l_pos by (auto simp: lift_profile_def) - -lemma size_lift_profile: "size (lift_profile A) = l * size A" - by (simp add: size_mset_sum_mset lift_profile_def image_mset.compositionality o_def) - -lemma count_lift_profile [simp]: "count (lift_profile A) x = l * count A x" - unfolding lift_profile_def by (induction A) auto - -lemma is_pref_profile_lift_profile [intro]: - assumes "lowered.is_pref_profile A" - shows "is_pref_profile (lift_profile A)" -proof - - interpret anon_papp_profile n_voters parties committee_size A - by fact - show ?thesis - using A_nonempty A_subset size_A - by unfold_locales - (auto simp: lift_profile_def size_mset_sum_mset image_mset.compositionality o_def) -qed - -sublocale lowered: anon_papp n_voters parties committee_size "r \ lift_profile" -proof - fix A assume "lowered.is_pref_profile A" - hence "is_pref_profile (lift_profile A)" - by blast - hence "is_committee (r (lift_profile A))" - using rule_wf by blast - thus "lowered.is_committee ((r \ lift_profile) A)" - by simp -qed - -sublocale lowered: weak_rep_anon_papp n_voters parties committee_size "r \ lift_profile" -proof - fix A x - assume A: "lowered.is_pref_profile A" and x: "n_voters \ committee_size * count A {x}" - from A have A': "is_pref_profile (lift_profile A)" - by blast - from x have "l * n_voters \ l * (committee_size * count A {x})" - by (rule mult_left_mono) auto - also have "\ = committee_size * count (lift_profile A) {x}" - by simp - finally have "x \# r (lift_profile A)" - by (intro weak_representation A') - thus "x \# (r \ lift_profile) A" - by simp -qed - -sublocale lowered: card_stratproof_anon_papp n_voters parties committee_size "r \ lift_profile" -proof - fix A X Y - show "\lowered.card_manipulable A X Y" - unfolding lowered.card_manipulable_def - proof (rule notI, elim conjE) - assume A: "lowered.is_pref_profile A" and XY: "X \# A" "Y \ {}" "Y \ parties" - assume *: "(r \ lift_profile) A \[lowered.committee_preference X] - (r \ lift_profile) (A - {#X#} + {#Y#})" - interpret anon_papp_profile n_voters parties committee_size A - by fact - have X: "X \ {}" "X \ parties" - using XY A_nonempty A_subset by auto - - define A' where "A' = A - {#X#}" - have A': "A = A' + {#X#}" - using XY by (simp add: A'_def) - - have "r (lift_profile A) \[committee_preference X] - r (lift_profile (A - {#X#} + {#Y#}))" - using * by simp - also have "r (lift_profile (A - {#X#} + {#Y#})) \[committee_preference X] - r (lift_profile A - {#X#} + {#Y#})" - proof - - have "r (replicate_mset (l - 1) Y + (lift_profile A' + {#Y#})) \[committee_preference X] - r (replicate_mset (l - 1) X + (lift_profile A' + {#Y#}))" - proof (rule proposition2) - show "size (lift_profile A' + {#Y#}) + (l - 1) = l * n_voters" - using XY l_pos n_voters_pos - by (simp add: A'_def size_lift_profile size_Diff_singleton - algebra_simps Suc_diff_le size_A) - next - show "{} \# lift_profile A' + {#Y#} + replicate_mset (l - 1) Y" - using XY A_nonempty by (auto simp: A'_def dest: in_diffD) - next - show "\X'\#lift_profile A' + {#Y#} + replicate_mset (l - 1) Y. X' \ parties" - using XY A_subset by (auto simp: A'_def dest: in_diffD) - qed (use X in auto) - thus ?thesis - by (simp add: A' replicate_mset_rec l_pos lift_profile_def) - qed - finally have "card_manipulable (lift_profile A) X Y" - unfolding card_manipulable_def using XY A by auto - with not_manipulable show False - by blast - qed -qed - -sublocale lowered: card_stratproof_weak_rep_anon_papp n_voters parties committee_size "r \ lift_profile" - .. - -end - - -locale divide_voters_card_stratproof_weak_prop_rep_anon_papp = - card_stratproof_weak_prop_rep_anon_papp "l * n_voters" parties committee_size r - for l n_voters parties committee_size r -begin - -sublocale divide_voters_card_stratproof_weak_rep_anon_papp .. - -sublocale lowered: card_stratproof_weak_prop_rep_anon_papp - n_voters parties committee_size "r \ lift_profile" -proof - fix A x l' - assume A: "lowered.is_pref_profile A" and x: "l' * n_voters \ committee_size * count A {x}" - from A have A': "is_pref_profile (lift_profile A)" - by blast - from x have "l * (l' * n_voters) \ l * (committee_size * count A {x})" - by (rule mult_left_mono) auto - also have "\ = committee_size * count (lift_profile A) {x}" - by simp - also have "l * (l' * n_voters) = l' * (l * n_voters)" - by (simp add: algebra_simps) - finally have "count (r (lift_profile A)) x \ l'" - by (intro weak_proportional_representation A') - thus "count ((r \ lift_profile) A) x \ l'" - by simp -qed - -end - - - -subsection \Decreasing the number of parties\ - -text \ - If we have a PAPP rule that satisfies weak representation and cardinality strategyproofness, - for $m$ parties, we can turn it into one for $m-1$ parties. This is done by simply - duplicating one particular party (say \x\) in the preference profile, i.e.\ whenever - \x\ is part of an approval list, we add a clone of \x\ (say \y\) as well. Should \y\ then end up - in the committee, we simply replace it with \x\. - - Consequently, if we have an impossibility result for $k$ parties, it also holds for - $\geq m$ parties. -\ -locale remove_alt_card_stratproof_weak_rep_anon_papp = - card_stratproof_weak_rep_anon_papp n_voters parties committee_size r - for n_voters and parties :: "'a set" and committee_size r + - fixes x y :: 'a - assumes xy: "x \ parties" "y \ parties" "x \ y" -begin - -definition lift_applist :: "'a set \ 'a set" where - "lift_applist X = (if x \ X then insert y X else X)" - -definition lift_profile :: "'a set multiset \ 'a set multiset" where - "lift_profile A = image_mset lift_applist A" - -definition lower_result where "lower_result C = image_mset (\z. if z = y then x else z) C" - -definition lowered where "lowered = lower_result \ r \ lift_profile" - - -lemma lift_profile_empty [simp]: "lift_profile {#} = {#}" - by (simp add: lift_profile_def) - -lemma lift_profile_add_mset [simp]: - "lift_profile (add_mset X A) = add_mset (lift_applist X) (lift_profile A)" - by (simp add: lift_profile_def) - -lemma empty_in_lift_profile_iff [simp]: "{} \# lift_profile A \ {} \# A" - by (auto simp: lift_applist_def lift_profile_def) - -lemma size_lift_profile [simp]: "size (lift_profile A) = size A" - by (simp add: lift_profile_def) - -lemma lift_applist_eq_self_iff [simp]: "lift_applist X = X \ x \ X \ y \ X" - by (auto simp: lift_applist_def) - -lemma lift_applist_eq_self_iff' [simp]: "lift_applist (X - {y}) = X \ (x \ X \ y \ X)" - by (cases "y \ X") (auto simp: lift_applist_def xy) - -lemma in_lift_applist_iff: "z \ lift_applist X \ z \ X \ (z = y \ x \ X)" - by (auto simp: lift_applist_def) - -lemma count_lift_profile: - assumes "\Y\#A. y \ Y" - shows "count (lift_profile A) X = (if x \ X \ y \ X then count A (X - {y}) else 0)" - using assms xy by (induction A) (auto simp: lift_applist_def) - -lemma y_notin_lower_result [simp]: "y \# lower_result C" - using xy by (auto simp: lower_result_def) - -lemma lower_result_subset: "set_mset (lower_result C) \ insert x (set_mset C - {y})" - using xy by (auto simp: lower_result_def) - -lemma lower_result_subset': "set_mset C \ parties \ set_mset (lower_result C) \ parties" - using xy by (auto simp: lower_result_def) - -lemma size_lower_result [simp]: "size (lower_result C) = size C" - by (simp add: lower_result_def) - -lemma count_lower_result: - "count (lower_result C) z = - (if z = y then 0 - else if z = x then count C x + count C y - else count C z)" - using xy by (induction C) (auto simp: lower_result_def) - -lemma in_lower_result_iff: - "z \# lower_result C \ z \ y \ (z \# C \ (z = x \ y \# C))" - unfolding lower_result_def using xy by (induction C) auto - - -sublocale lowered: anon_papp_election n_voters "parties - {y}" - by standard (use n_voters_pos xy in auto) - -lemma is_pref_profile_lift_profile [intro]: - assumes "lowered.is_pref_profile A" - shows "is_pref_profile (lift_profile A)" -proof - - interpret anon_papp_profile n_voters "parties - {y}" committee_size A - by fact - show ?thesis - using A_nonempty A_subset size_A - by unfold_locales - (auto simp: lift_profile_def lift_applist_def xy - size_mset_sum_mset image_mset.compositionality o_def) -qed - -sublocale lowered: anon_papp n_voters "parties - {y}" committee_size lowered -proof - fix A assume "lowered.is_pref_profile A" - hence "is_pref_profile (lift_profile A)" - by blast - hence "is_committee (r (lift_profile A))" - using rule_wf by blast - thus "lowered.is_committee (lowered A)" - unfolding lowered.is_committee_def is_committee_def lowered_def - using lower_result_subset'[of "r (lift_profile A)"] by auto -qed - -sublocale lowered: weak_rep_anon_papp n_voters "parties - {y}" committee_size lowered -proof - fix A z - assume A: "lowered.is_pref_profile A" and z: "n_voters \ committee_size * count A {z}" - interpret A: anon_papp_profile n_voters "parties - {y}" committee_size A - by fact - have "committee_size > 0" - using z n_voters_pos by (intro Nat.gr0I) auto - - from A have A': "is_pref_profile (lift_profile A)" - by blast - have "count A {z} > 0" - using z n_voters_pos by (intro Nat.gr0I) auto - hence "{z} \# A" - by simp - hence z': "z \ parties - {y}" - using A.A_subset z by auto - - define C where "C = r (lift_profile A)" - - show "z \# lowered A" - proof (cases "z = x") - case False - have "n_voters \ committee_size * count A {z}" - by fact - also have "count A {z} \ count (lift_profile A) {z}" - using A.A_subset z' False by (subst count_lift_profile) auto - hence "committee_size * count A {z} \ committee_size * count (lift_profile A) {z}" - by (intro mult_left_mono) auto - finally have "z \# r (lift_profile A)" - by (intro weak_representation A') - thus "z \# lowered A" - using False z' by (simp add: lowered_def in_lower_result_iff) - next - case [simp]: True - have "1 \ size {#z \# C. z \ {x, y}#}" - unfolding C_def - proof (rule proposition3) - have [simp]: "{x, y} - {y} = {x}" - using xy by auto - hence "n_voters \ committee_size * count (lift_profile A) {x, y}" - using xy A.A_subset z by (subst count_lift_profile) auto - thus "int 1 * \real n_voters / real committee_size\ \ int (count (lift_profile A) {x, y})" - using \committee_size > 0\ - by (auto simp: ceiling_le_iff field_simps simp flip: of_nat_mult) - qed (use A' xy \committee_size > 0\ in auto) - also have "\ = count C x + count C y" - using xy by (induction C) auto - also have "\ = count (lowered A) x" - using xy by (simp add: lowered_def count_lower_result C_def) - finally show "z \# lowered A" - by simp - qed -qed - -lemma filter_lower_result_eq: - "y \ X \ {#x \# lower_result C. x \ X#} = lower_result {#x \# C. x \ lift_applist X#}" - by (induction C) (auto simp: lower_result_def lift_applist_def) - -sublocale lowered: card_stratproof_anon_papp n_voters "parties - {y}" committee_size lowered -proof - fix A X Y - show "\lowered.card_manipulable A X Y" - unfolding lowered.card_manipulable_def - proof (rule notI, elim conjE) - assume A: "lowered.is_pref_profile A" and XY: "X \# A" "Y \ {}" "Y \ parties - {y}" - assume *: "lowered A \[lowered.committee_preference X] lowered (A - {#X#} + {#Y#})" - interpret anon_papp_profile n_voters "parties - {y}" committee_size A - by fact - have X: "X \ {}" "X \ parties - {y}" - using XY A_nonempty A_subset by auto - define A' where "A' = A - {#X#}" - have A': "A = A' + {#X#}" - using XY by (simp add: A'_def) - - from * have "size {#x \# lower_result (r (lift_profile A' + {#lift_applist X#})). x \ X#} < - size {#x \# lower_result (r (lift_profile A' + {#lift_applist Y#})). x \ X#}" - by (simp add: lowered_def A' lowered.strong_committee_preference_iff) - also have "{#x \# lower_result (r (lift_profile A' + {#lift_applist X#})). x \ X#} = - lower_result {#x \# r (lift_profile A' + {#lift_applist X#}). x \ lift_applist X#}" - using X by (subst filter_lower_result_eq) auto - also have "{#x \# lower_result (r (lift_profile A' + {#lift_applist Y#})). x \ X#} = - lower_result {#x \# r (lift_profile A' + {#lift_applist Y#}). x \ lift_applist X#}" - using X by (subst filter_lower_result_eq) auto - finally have "size {#x \# r (lift_profile A' + {#lift_applist X#}). x \ lift_applist X#} < - size {#x \# r (lift_profile A' + {#lift_applist Y#}). x \ lift_applist X#}" - by simp - hence "r (lift_profile A' + {#lift_applist X#}) \[committee_preference (lift_applist X)] - r (lift_profile A' + {#lift_applist Y#})" - by (simp add: strong_committee_preference_iff) - moreover have "\r (lift_profile A' + {#lift_applist X#}) \[committee_preference (lift_applist X)] - r (lift_profile A' + {#lift_applist Y#})" - proof (rule not_manipulable' [where Y = "lift_applist Y"]) - have "is_pref_profile (lift_profile A)" - using A by blast - thus "is_pref_profile (lift_profile A' + {#lift_applist X#})" - using A by (simp add: A') - next - have "is_pref_profile (lift_profile (A - {#X#} + {#Y#}))" - using A XY lowered.is_pref_profile_replace by blast - thus "is_pref_profile (lift_profile A' + {#lift_applist Y#})" - by (simp add: A') - qed auto - ultimately show False - by contradiction - qed -qed - -sublocale lowered: card_stratproof_weak_rep_anon_papp n_voters "parties - {y}" committee_size lowered - .. - -end - - -text \ - The following lemma is now simply an iterated application of the above. This allows us to - restrict a P-APP rule to any non-empty subset of parties. -\ -lemma card_stratproof_weak_rep_anon_papp_restrict_parties: - assumes "card_stratproof_weak_rep_anon_papp n parties k r" "parties' \ parties" "parties' \ {}" - shows "\r. card_stratproof_weak_rep_anon_papp n parties' k r" -proof - - have "finite parties" - proof - - interpret card_stratproof_weak_rep_anon_papp n parties k r - by fact - show ?thesis - by (rule finite_parties) - qed - thus ?thesis - using assms - proof (induction parties arbitrary: r rule: finite_psubset_induct) - case (psubset parties r) - show ?thesis - proof (cases "parties = parties'") - case True - thus ?thesis - using psubset.prems by auto - next - case False - obtain x where x: "x \ parties'" - using psubset.prems by blast - from False and psubset.prems have "parties - parties' \ {}" - by auto - then obtain y where y: "y \ parties - parties'" - by blast - - interpret card_stratproof_weak_rep_anon_papp n parties k r - by fact - interpret remove_alt_card_stratproof_weak_rep_anon_papp n parties k r x y - by standard (use x y psubset.prems in auto) - show ?thesis - proof (rule psubset.IH) - show "parties - {y} \ parties" and "parties' \ parties - {y}" "parties' \ {}" - using x y psubset.prems by auto - next - show "card_stratproof_weak_rep_anon_papp n (parties - {y}) k lowered" - using lowered.card_stratproof_weak_rep_anon_papp_axioms . - qed - qed - qed -qed - - - -subsection \Decreasing the committee size\ - -text \ - If we have a PAPP rule that satisfies weak representation and cardinality strategyproofness, - for $l(k+1)$ voters, $m+1$ parties, and a committee size of $k+1$, we can turn it into - one for $lk$ voters, $m$ parties, and a committee size of $k$. - - This is done by again cloning a party \x\ into a new party \y\ and additionally - adding \l\ new voters whose preferences are \{x, y}\. We again replace any \y\ occuring in the - output committee with \x\. Weak representation then ensures that \x\ occurs in the output - at least once, and we then simply remove one \x\ from the committee to obtain an output - committee of size \k - 1\. - - Consequently, if we have an impossibility result for a committee size of $m$, we can - extend it to a larger committee size, but at the cost of introducing a new party - and new voters, and with a restriction on the number of voters. -\ -(* n / k = l *) -locale decrease_committee_card_stratproof_weak_rep_anon_papp = - card_stratproof_weak_rep_anon_papp "l * (k + 1)" "insert y parties" "k + 1" r - for l k y and parties :: "'a set" and r + - fixes x :: 'a - assumes xy: "x \ parties" "y \ parties" - assumes k: "k > 0" -begin - -definition lift_applist :: "'a set \ 'a set" where - "lift_applist X = (if x \ X then insert y X else X)" - -definition lift_profile :: "'a set multiset \ 'a set multiset" where - "lift_profile A = image_mset lift_applist A + replicate_mset l {x, y}" - -definition lower_result - where "lower_result C = image_mset (\z. if z = y then x else z) C - {#x#}" - -definition lowered where "lowered = lower_result \ r \ lift_profile" - - -lemma l: "l > 0" - using n_voters_pos by auto - -lemma x_neq_y [simp]: "x \ y" "y \ x" - using xy by auto - -lemma lift_profile_empty [simp]: "lift_profile {#} = replicate_mset l {x, y}" - by (simp add: lift_profile_def) - -lemma lift_profile_add_mset [simp]: - "lift_profile (add_mset X A) = add_mset (lift_applist X) (lift_profile A)" - by (simp add: lift_profile_def) - -lemma empty_in_lift_profile_iff [simp]: "{} \# lift_profile A \ {} \# A" - by (auto simp: lift_applist_def lift_profile_def) - -lemma size_lift_profile [simp]: "size (lift_profile A) = size A + l" - by (simp add: lift_profile_def) - -lemma lift_applist_eq_self_iff [simp]: "lift_applist X = X \ x \ X \ y \ X" - by (auto simp: lift_applist_def) - -lemma lift_applist_eq_self_iff' [simp]: "lift_applist (X - {y}) = X \ (x \ X \ y \ X)" - by (cases "y \ X") (auto simp: lift_applist_def xy) - -lemma in_lift_applist_iff: "z \ lift_applist X \ z \ X \ (z = y \ x \ X)" - by (auto simp: lift_applist_def) - -lemma count_lift_profile: - assumes "\Y\#A. y \ Y" - shows "count (lift_profile A) X = - (if x \ X \ y \ X then count A (X - {y}) else 0) + - (if X = {x, y} then l else 0)" - using assms xy by (induction A) (auto simp: lift_applist_def) - - - -lemma y_notin_lower_result [simp]: "y \# lower_result C" - using xy by (auto simp: lower_result_def dest: in_diffD) - -lemma lower_result_subset: "set_mset (lower_result C) \ insert x (set_mset C - {y})" - using xy by (auto simp: lower_result_def dest: in_diffD) - -lemma lower_result_subset': "set_mset C \ insert y parties \ set_mset (lower_result C) \ parties" - by (rule order.trans[OF lower_result_subset]) (use xy in auto) - -lemma size_lower_result [simp]: - assumes "x \# C \ y \# C" - shows "size (lower_result C) = size C - 1" - using assms by (auto simp: lower_result_def size_Diff_singleton) - -lemma size_lower_result': "size (lower_result C) = size C - (if x \# C \ y \# C then 1 else 0)" -proof - - define f where "f = (\C. image_mset (\z. if z = y then x else z) C)" - have "size (lower_result C) = size (f C - {#x#})" - by (simp add: lower_result_def f_def) - also have "\ = size (f C) - (if x \# f C then 1 else 0)" - by (simp add: diff_single_trivial size_Diff_singleton) - finally show ?thesis - by (auto simp: f_def) -qed - -lemma count_lower_result: - "count (lower_result C) z = - (if z = y then 0 - else if z = x then count C x + count C y - 1 - else count C z)" (is "_ = ?rhs") -proof - - define f where "f = (\C. image_mset (\z. if z = y then x else z) C)" - have "count (lower_result C) z = count (f C - {#x#}) z" - by (simp add: lower_result_def f_def) - also have "\ = count (f C) z - (if z = x then 1 else 0)" - by auto - also have "count (f C) z = (if z = y then 0 else if z = x then count C x + count C y else count C z)" - using xy by (induction C) (auto simp: f_def) - also have "\ - (if z = x then 1 else 0) = ?rhs" - by auto - finally show ?thesis . -qed - -lemma in_lower_resultD: - "z \# lower_result C \ z = x \ z \# C" - using xy by (auto simp: lower_result_def dest!: in_diffD) - -lemma in_lower_result_iff: - "z \# lower_result C \ z \ y \ (if z = x then count C x + count C y > 1 else z \# C)" - (is "_ = ?rhs") -proof - - have "z \# lower_result C \ count (lower_result C) z > 0" - by auto - also have "\ \ ?rhs" - by (subst count_lower_result) auto - finally show ?thesis . -qed - -lemma filter_lower_result_eq: - assumes "y \ X" - shows "{#z \# lower_result C. z \ X#} = lower_result {#z \# C. z \ lift_applist X#}" -proof - - define f where "f = (\C. {#if z = y then x else z. z \# C#})" - have "lower_result {#z \# C. z \ lift_applist X#} = f {#z \# C. z \ lift_applist X#} - {#x#}" - by (simp add: f_def lower_result_def) - also have "f {#z \# C. z \ lift_applist X#} = {#z \# f C. z \ X#}" - using assms by (induction C) (auto simp: f_def lift_applist_def) - also have "\ - {#x#} = {#z \# f C - {#x#}. z \ X#}" - by (subst filter_diff_mset') auto - also have "f C - {#x#} = lower_result C" - by (simp add: f_def lower_result_def) - finally show ?thesis .. -qed - - -sublocale lowered: anon_papp_election "l * k" parties k - by standard (use n_voters_pos xy finite_parties k in auto) - -lemma is_pref_profile_lift_profile [intro]: - assumes "lowered.is_pref_profile A" - shows "is_pref_profile (lift_profile A)" -proof - - interpret A: anon_papp_profile "l * k" parties k A - by fact - show ?thesis - using A.A_nonempty A.A_subset A.size_A l - by unfold_locales - (auto simp: lift_profile_def lift_applist_def xy - size_mset_sum_mset image_mset.compositionality o_def) -qed - -lemma is_committee_lower_result: - assumes "is_committee C" "x \# C \ y \# C" - shows "lowered.is_committee (lower_result C)" - using assms unfolding is_committee_def lowered.is_committee_def - using lower_result_subset'[of C] by auto - -lemma x_or_y_in_r_lift_profile: - assumes "lowered.is_pref_profile A" - shows "x \# r (lift_profile A) \ y \# r (lift_profile A)" -proof - - interpret A: anon_papp_profile "l * k" parties k A - by fact - have "size {#z \# r (lift_profile A). z \ {x, y}#} \ 1" - proof (rule proposition3) - have "real (l * (k + 1)) / real (k + 1) = real l" - by (simp add: field_simps) - also have "int 1 * \\\ = int l" - by simp - also have "l \ count (lift_profile A) {x, y}" - using xy A.A_subset by (subst count_lift_profile) auto - finally show "int 1 * \real (l * (k + 1)) / real (k + 1)\ \ int (count (lift_profile A) {x, y})" - by simp - next - show "is_pref_profile (lift_profile A)" - by (intro is_pref_profile_lift_profile) fact - qed (use xy in auto) - hence "{#z \# r (lift_profile A). z \ {x, y}#} \ {#}" - by force - thus ?thesis - by auto -qed - -sublocale lowered: anon_papp "l * k" parties k lowered -proof - fix A assume A: "lowered.is_pref_profile A" - hence "is_pref_profile (lift_profile A)" - by blast - hence "is_committee (r (lift_profile A))" - using rule_wf by blast - thus "lowered.is_committee (lowered A)" - unfolding lowered_def o_def using x_or_y_in_r_lift_profile[of A] A - by (intro is_committee_lower_result) auto -qed - -sublocale lowered: weak_rep_anon_papp "l * k" parties k lowered -proof - fix A z - assume A: "lowered.is_pref_profile A" and z: "l * k \ k * count A {z}" - interpret A: anon_papp_profile "l * k" "parties" k A - by fact - - from A have A': "is_pref_profile (lift_profile A)" - by blast - have "count A {z} > 0" - using z k n_voters_pos by (intro Nat.gr0I) auto - hence "{z} \# A" - by simp - hence z': "z \ parties" - using A.A_subset z by auto - hence [simp]: "y \ z" "z \ y" - using xy by auto - - define C where "C = r (lift_profile A)" - - show "z \# lowered A" - proof (cases "z = x") - case False - have "l \ count A {z}" - using z k by (simp add: algebra_simps) - hence "l * (k + 1) \ (k + 1) * count A {z}" - by (subst mult.commute, intro mult_right_mono) auto - also have "count A {z} = count (lift_profile A) {z}" - using False A.A_subset xy by (subst count_lift_profile) auto - finally have "z \# r (lift_profile A)" - by (intro weak_representation A') - thus "z \# lowered A" - using False by (auto simp: lowered_def in_lower_result_iff) - next - case [simp]: True - from xy have [simp]: "{x, y} - {y} = {x}" - by auto - - have "size {#z \# C. z \ {x, y}#} \ 2" - unfolding C_def - proof (rule proposition3) - have "real (l * (k + 1)) / real (k + 1) = l" - unfolding of_nat_mult using k by (simp add: divide_simps) - also have "int 2 * \\\ = int (2 * l)" - by simp - also have "\ \ count (lift_profile A) {x, y}" - using z k xy A.A_subset by (subst count_lift_profile) auto - finally show "int 2 * \real (l * (k + 1)) / real (k + 1)\ \ \" . - qed (use A' xy in auto) - also have "size {#z \# C. z \ {x, y}#} = count C x + count C y" - by (induction C) auto - finally have "x \# lower_result C" - by (subst in_lower_result_iff) auto - thus "z \# lowered A" - by (simp add: lowered_def C_def) - qed -qed - -sublocale lowered: card_stratproof_anon_papp "l * k" parties k lowered -proof - fix A X Y - show "\lowered.card_manipulable A X Y" - unfolding lowered.card_manipulable_def - proof (rule notI, elim conjE) - assume A: "lowered.is_pref_profile A" and XY: "X \# A" "Y \ {}" "Y \ parties" - assume *: "lowered A \[lowered.committee_preference X] lowered (A - {#X#} + {#Y#})" - interpret anon_papp_profile "l * k" "parties" k A - by fact - have X: "X \ {}" "X \ parties" - using XY A_nonempty A_subset by auto - define A' where "A' = A - {#X#}" - have A': "A = A' + {#X#}" - using XY by (simp add: A'_def) - from xy X XY have [simp]: "y \ X" "y \ Y" - by auto - - define Al1 where "Al1 = lift_profile A" - define Al2 where "Al2 = lift_profile (A' + {#Y#})" - have A'_plus_Y: "lowered.is_pref_profile (A' + {#Y#})" - unfolding A'_def using A XY lowered.is_pref_profile_replace by blast - have Al1: "is_pref_profile Al1" - unfolding Al1_def using A by blast - have Al2: "is_pref_profile Al2" - unfolding Al2_def unfolding A'_def using A XY lowered.is_pref_profile_replace by blast - - have size_aux: "size (lower_result {#x \# r (lift_profile A). x \ lift_applist X#}) = - size {#x \# r (lift_profile A). x \ lift_applist X#} - (if x \ X then 1 else 0)" - if A: "lowered.is_pref_profile A" for A - using x_or_y_in_r_lift_profile[OF A] - by (subst size_lower_result') (auto simp: lift_applist_def) - - from * have "size {#x \# lower_result (r Al1). x \ X#} < - size {#x \# lower_result (r Al2). x \ X#}" - by (simp add: Al1_def Al2_def lowered_def A' lowered.strong_committee_preference_iff) - also have "{#x \# lower_result (r Al1). x \ X#} = lower_result {#x \# r Al1. x \ lift_applist X#}" - using X xy by (subst filter_lower_result_eq) auto - also have "{#x \# lower_result (r Al2). x \ X#} = lower_result {#x \# r Al2. x \ lift_applist X#}" - using X xy by (subst filter_lower_result_eq) auto - also have "size (lower_result {#x \# r Al1. x \ lift_applist X#}) = - size {#x \# r Al1. x \ lift_applist X#} - (if x \ X then 1 else 0)" - unfolding Al1_def by (rule size_aux) fact - also have "size (lower_result {#x \# r Al2. x \ lift_applist X#}) = - size {#x \# r Al2. x \ lift_applist X#} - (if x \ X then 1 else 0)" - unfolding Al2_def by (rule size_aux) fact - finally have "size {#x \# r Al1. x \ lift_applist X#} < size {#x \# r Al2. x \ lift_applist X#}" - by auto - hence "r Al1 \[committee_preference (lift_applist X)] r Al2" - by (simp add: strong_committee_preference_iff) - - moreover have "\r Al1 \[committee_preference (lift_applist X)] r Al2" - by (rule not_manipulable' [where Y = "lift_applist Y"]) - (use Al1 Al2 in \auto simp: Al1_def Al2_def A'\) - - ultimately show False - by contradiction - qed -qed - -sublocale lowered: card_stratproof_weak_rep_anon_papp "l * k" parties k lowered - .. - -end - - -text \ - For Weak \<^emph>\Proportional\ Representation, the lowering argument to decrease the committee size - is somewhat easier since it does not involve adding a new party; instead, we simply - add \l\ new voters whose preferences are \{x}\. -\ -locale decrease_committee_card_stratproof_weak_prop_rep_anon_papp = - card_stratproof_weak_prop_rep_anon_papp "l * (k + 1)" parties "k + 1" r - for l k and parties :: "'a set" and r + - fixes x :: 'a - assumes x: "x \ parties" - assumes k: "k > 0" -begin - -definition lift_profile :: "'a set multiset \ 'a set multiset" where - "lift_profile A = A + replicate_mset l {x}" - -definition lower_result - where "lower_result C = C - {#x#}" - -definition lowered where "lowered = lower_result \ r \ lift_profile" - - -lemma l: "l > 0" - using n_voters_pos by auto - -lemma lift_profile_empty [simp]: "lift_profile {#} = replicate_mset l {x}" - by (simp add: lift_profile_def) - -lemma lift_profile_add_mset [simp]: - "lift_profile (add_mset X A) = add_mset X (lift_profile A)" - by (simp add: lift_profile_def) - -lemma empty_in_lift_profile_iff [simp]: "{} \# lift_profile A \ {} \# A" - by (auto simp: lift_profile_def) - -lemma size_lift_profile [simp]: "size (lift_profile A) = size A + l" - by (simp add: lift_profile_def) - -lemma count_lift_profile: - "count (lift_profile A) X = count A X + (if X = {x} then l else 0)" - by (auto simp: lift_profile_def) - - -lemma size_lower_result [simp]: - assumes "x \# C" - shows "size (lower_result C) = size C - 1" - using assms by (auto simp: lower_result_def size_Diff_singleton) - -lemma size_lower_result': "size (lower_result C) = size C - (if x \# C then 1 else 0)" - by (induction C) (auto simp: lower_result_def size_Diff_singleton) - -lemma count_lower_result: - "count (lower_result C) z = count C z - (if z = x then 1 else 0)" - by (auto simp: lower_result_def) - -lemma in_lower_resultD: - "z \# lower_result C \ z \# C" - by (auto simp: lower_result_def dest!: in_diffD) - -lemma in_lower_result_iff: - "z \# lower_result C \ (if z = x then count C x > 1 else z \# C)" - (is "_ = ?rhs") -proof - - have "z \# lower_result C \ count (lower_result C) z > 0" - by auto - also have "\ \ ?rhs" - by (subst count_lower_result) auto - finally show ?thesis . -qed - - -sublocale lowered: anon_papp_election "l * k" parties k - by standard (use n_voters_pos finite_parties k in auto) - -lemma is_pref_profile_lift_profile [intro]: - assumes "lowered.is_pref_profile A" - shows "is_pref_profile (lift_profile A)" -proof - - interpret A: anon_papp_profile "l * k" parties k A - by fact - show ?thesis - using A.A_nonempty A.A_subset A.size_A l - by unfold_locales - (auto simp: lift_profile_def x size_mset_sum_mset image_mset.compositionality o_def) -qed - -lemma is_committee_lower_result: - assumes "is_committee C" "x \# C" - shows "lowered.is_committee (lower_result C)" - using assms unfolding is_committee_def lowered.is_committee_def - by (auto simp: lower_result_def size_Diff_singleton dest: in_diffD) - -lemma x_in_r_lift_profile: - assumes "lowered.is_pref_profile A" - shows "x \# r (lift_profile A)" -proof (rule weak_representation) - show "is_pref_profile (lift_profile A)" - using assms by blast -next - have "(k + 1) * l \ (k + 1) * count (lift_profile A) {x}" - by (intro mult_left_mono) (auto simp: count_lift_profile) - thus "l * (k + 1) \ (k + 1) * count (lift_profile A) {x}" - by (simp add: algebra_simps) -qed - -sublocale lowered: anon_papp "l * k" parties k lowered -proof - fix A assume A: "lowered.is_pref_profile A" - hence "is_pref_profile (lift_profile A)" - by blast - hence "is_committee (r (lift_profile A))" - using rule_wf by blast - thus "lowered.is_committee (lowered A)" - unfolding lowered_def o_def using x_in_r_lift_profile[of A] A - by (intro is_committee_lower_result) auto -qed - -sublocale lowered: weak_prop_rep_anon_papp "l * k" parties k lowered -proof - fix A z l' - assume A: "lowered.is_pref_profile A" and z: "l' * (l * k) \ k * count A {z}" - interpret A: anon_papp_profile "l * k" "parties" k A - by fact - - show "count (lowered A) z \ l'" - proof (cases "l' > 0") - case False - thus ?thesis by auto - next - case l: True - - from A have A': "is_pref_profile (lift_profile A)" - by blast - have "count A {z} > 0" - using z k n_voters_pos l by (intro Nat.gr0I) auto - hence "{z} \# A" - by simp - hence z': "z \ parties" - using A.A_subset z by auto - - define C where "C = r (lift_profile A)" - - show "count (lowered A) z \ l'" - proof (cases "z = x") - case False - have "l' * l \ count A {z}" - using z k by (simp add: algebra_simps) - hence "l' * l * (k + 1) \ (k + 1) * count A {z}" - by (subst mult.commute, intro mult_right_mono) auto - also have "count A {z} = count (lift_profile A) {z}" - using False A.A_subset by (subst count_lift_profile) auto - finally have "count (r (lift_profile A)) z \ l'" - by (intro weak_proportional_representation A') (auto simp: algebra_simps) - thus "l' \ count (lowered A) z" - using False by (simp add: lowered_def lower_result_def) - next - case [simp]: True - have "l' * l \ count A {x}" - using z k by (simp add: algebra_simps) - hence "l' * l * (k + 1) \ (k + 1) * count A {x}" - by (subst mult.commute, intro mult_right_mono) auto - also have "\ + (k + 1) * l = (k + 1) * count (lift_profile A) {x}" - by (simp add: count_lift_profile algebra_simps) - finally have "(l' + 1) * (l * (k + 1)) \ (k + 1) * count (lift_profile A) {x}" - by (simp add: algebra_simps) - hence "count (r (lift_profile A)) x \ l' + 1" - by (intro weak_proportional_representation A') - thus "l' \ count (lowered A) z" - by (simp add: lowered_def lower_result_def) - qed - qed -qed - -sublocale lowered: card_stratproof_anon_papp "l * k" parties k lowered -proof - fix A X Y - show "\lowered.card_manipulable A X Y" - unfolding lowered.card_manipulable_def - proof (rule notI, elim conjE) - assume A: "lowered.is_pref_profile A" and XY: "X \# A" "Y \ {}" "Y \ parties" - assume *: "lowered A \[lowered.committee_preference X] lowered (A - {#X#} + {#Y#})" - interpret anon_papp_profile "l * k" "parties" k A - by fact - have X: "X \ {}" "X \ parties" - using XY A_nonempty A_subset by auto - define A' where "A' = A - {#X#}" - have A': "A = A' + {#X#}" - using XY by (simp add: A'_def) - - define Al1 where "Al1 = lift_profile A" - define Al2 where "Al2 = lift_profile (A' + {#Y#})" - have A'_plus_Y: "lowered.is_pref_profile (A' + {#Y#})" - unfolding A'_def using A XY lowered.is_pref_profile_replace by blast - have Al1: "is_pref_profile Al1" - unfolding Al1_def using A by blast - have Al2: "is_pref_profile Al2" - unfolding Al2_def unfolding A'_def using A XY lowered.is_pref_profile_replace by blast - - have size_aux: "size (lower_result {#x \# r (lift_profile A). x \ X#}) = - size {#x \# r (lift_profile A). x \ X#} - (if x \ X then 1 else 0)" - if A: "lowered.is_pref_profile A" for A - using x_in_r_lift_profile[OF A] by (subst size_lower_result') auto - - from * have "size {#x \# lower_result (r Al1). x \ X#} < - size {#x \# lower_result (r Al2). x \ X#}" - by (simp add: Al1_def Al2_def lowered_def A' lowered.strong_committee_preference_iff) - also have "{#x \# lower_result (r Al1). x \ X#} = lower_result {#x \# r Al1. x \ X#}" - using X x unfolding lower_result_def by (subst filter_diff_mset') auto - also have "{#x \# lower_result (r Al2). x \ X#} = lower_result {#x \# r Al2. x \ X#}" - using X x unfolding lower_result_def by (subst filter_diff_mset') auto - also have "size (lower_result {#x \# r Al1. x \ X#}) = - size {#x \# r Al1. x \ X#} - (if x \ X then 1 else 0)" - unfolding Al1_def by (rule size_aux) fact - also have "size (lower_result {#x \# r Al2. x \ X#}) = - size {#x \# r Al2. x \ X#} - (if x \ X then 1 else 0)" - unfolding Al2_def by (rule size_aux) fact - finally have "size {#x \# r Al1. x \ X#} < size {#x \# r Al2. x \ X#}" - by auto - hence "r Al1 \[committee_preference X] r Al2" - by (simp add: strong_committee_preference_iff) - - moreover have "\r Al1 \[committee_preference X] r Al2" - by (rule not_manipulable' [where Y = "Y"]) - (use Al1 Al2 in \auto simp: Al1_def Al2_def A'\) - - ultimately show False - by contradiction - qed -qed - -sublocale lowered: card_stratproof_weak_prop_rep_anon_papp "l * k" parties k lowered - .. - -end - -end diff --git a/thys/PAPP_Impossibility/#ROOT# b/thys/PAPP_Impossibility/#ROOT# deleted file mode 100644 --- a/thys/PAPP_Impossibility/#ROOT# +++ /dev/null @@ -1,14 +0,0 @@ -chapter AFP - -session PAPP_Impossibility (AFP) = "HOL-Library" + - options [timeout = 900] - sessions - Randomised_Social_Choice - theories [document=false] - SAT_Replay - theories - PAPP_Impossibility - document_files - root.tex - root.bib - diff --git a/thys/PAPP_Impossibility/Anonymous_PAPP_Lowering.thy b/thys/PAPP_Impossibility/Anonymous_PAPP_Lowering.thy --- a/thys/PAPP_Impossibility/Anonymous_PAPP_Lowering.thy +++ b/thys/PAPP_Impossibility/Anonymous_PAPP_Lowering.thy @@ -1,1374 +1,1309 @@ (* File: Anonymous_PAPP_Lowering.thy Author: Manuel Eberl, University of Innsbruck *) section \Lowering P-APP Rules to Smaller Settings\ theory Anonymous_PAPP_Lowering imports Anonymous_PAPP begin text \ In this section, we prove a number of lemmas (corresponding to Lemma~1 in the paper) that allow us to take an anonymous P-APP rule with some additional properties (typically Cardinality-Strategyproofness and Weak Representation or Weak Proportional Representation) and construct from it an anonymous P-APP rule for a different setting, i.e.\ different number of voters, parties, and/or result committee size. In the reverse direction, this also allows us to lift impossibility results from one setting to another. \ subsection \Preliminary Lemmas\ context card_stratproof_anon_papp begin text \ The following lemma is obtained by applying Strategyproofness repeatedly. It shows that if we have \l\ voters with identical approval lists, then this entire group of voters has no incentive to submit wrong preferences. That is, the outcome they obtain by submitting their genuine approval lists is weakly preferred by them over all outcomes obtained where these \l\ voters submit any other preferences (and the remaining $n - l$ voters submit the same preferences as before). This is stronger than regular Strategyproofness, where we only demand that no voter has an incentive to submit wrong preferences \<^emph>\unilaterally\ (and everyone else keeps the same preferences). Here we know that the entire group of \l\ voters has no incentive to submit wrong preferences in coordination with one another. \ lemma proposition2: assumes "size B = l" "size A + l = n_voters" assumes "X \ {}" "X \ parties" "{} \# A+B" "\X'\#A+B. X' \ parties" shows "r (replicate_mset l X + A) \[Comm(X)] r (B + A)" using assms proof (induction l arbitrary: A B) case 0 thus ?case by simp next case (Suc l A B) from Suc.prems have "set_mset B \ {}" by auto then obtain Y where Y: "Y \# B" by blast define B' where "B' = B - {#Y#}" define A' where "A' = A + {#Y#}" have [simp]: "size B' = l" using Suc.prems Y by (simp add: B'_def size_Diff_singleton) have [simp]: "size A' = n_voters - l" using Suc.prems Y by (simp add: A'_def) have "r (B' + A') \[Comm(X)] r (replicate_mset l X + A')" by (rule Suc.IH) (use Suc.prems Y in \auto simp: A'_def B'_def size_Diff_singleton\) also have "B' + A' = B + A" using Y by (simp add: B'_def A'_def) also have "r (replicate_mset l X + A') \[Comm(X)] r (replicate_mset (Suc l) X + A)" proof (rule not_manipulable'') show "replicate_mset (Suc l) X + A + {#Y#} = replicate_mset l X + A' + {#X#}" by (simp add: A'_def) next show "is_pref_profile (replicate_mset (Suc l) X + A)" using Suc.prems by unfold_locales (auto split: if_splits) next show "is_pref_profile (replicate_mset l X + A')" using Suc.prems Y by unfold_locales (auto split: if_splits simp: A'_def) qed finally show ?case . qed end context card_stratproof_weak_rep_anon_papp begin text \ In a setting with Weak Representation and Cardinality-Strategyproofness, Proposition 2 allows us to strengthen Weak Representation in the following way: Suppose we at least $l \lfloor n/k\rfloor$ voters with the same approval list \X\, and \X\ consists of at least \l\ parties. Then at least \l\ of the members of the result committee are in \X\. \ lemma proposition3: assumes "is_pref_profile A" "X \ parties" "card X \ l" assumes "committee_size > 0" assumes "count A X \ l * \n_voters / committee_size\" shows "size {# x\# r A. x \ X #} \ l" using assms proof (induction l arbitrary: A X rule: less_induct) case (less l A X) interpret A: anon_papp_profile n_voters parties committee_size A by fact consider "l = 0" | "l = 1" | "l > 1" by force thus ?case proof cases assume "l = 0" thus ?thesis by simp next assume [simp]: "l = 1" define n where "n = count A X" with less.prems have "X \ {}" by auto then obtain x where x: "x \ X" by blast have "n \ size A" unfolding n_def by (rule count_le_size) hence "n \ n_voters" by (simp add: A.size_A) have "count A X > 0" by (rule Nat.gr0I) (use n_voters_pos less.prems in \auto simp: field_simps\) hence "X \# A" by force have [simp]: "replicate_mset n X \# A" by (simp add: n_def flip: count_le_replicate_mset_subset_eq) define A'' where "A'' = A - replicate_mset n X" define A' where "A' = A'' + replicate_mset n {x}" interpret A': anon_papp_profile n_voters parties committee_size A' using A.A_nonempty A.A_subset A.size_A x \X \# A\ by unfold_locales (fastforce simp: A'_def A''_def size_Diff_submset subset_mset.add_increasing2 split: if_splits dest!: in_diffD)+ have "x \# r A'" proof (rule weak_representation) show "is_pref_profile A'" by (fact A'.anon_papp_profile_axioms) next have "n_voters \ committee_size * n" using less.prems by (simp add: n_def ceiling_le_iff field_simps flip: of_nat_mult) also have "n \ count A' {x}" by (simp add: A'_def) finally show "n_voters \ committee_size * count A' {x}" by simp qed hence "1 \ count (r A') x" by simp also have "\ = size {# y \# r A'. y = x #}" by simp also have "\ \ size {# y \# r A'. y \ X #}" by (intro size_mset_mono multiset_filter_mono') (use x in auto) also have "r A' \[Comm(X)] r A" proof - have "r (replicate_mset n {x} + A'') \[Comm(X)] r (replicate_mset n X + A'')" proof (rule proposition2) show "{} \# A'' + replicate_mset n {x}" using A'.A_nonempty by (auto simp: A'_def) show "\X'\#A'' + replicate_mset n {x}. X' \ parties" using A'.A_subset x by (auto simp: A'_def dest: in_diffD) show "size A'' + n = n_voters" using \n \ n_voters\ by (auto simp: A''_def size_Diff_submset A.size_A) qed (use less.prems in auto) also have "replicate_mset n X + A'' = A" by (simp add: A''_def n_def flip: count_le_replicate_mset_subset_eq) finally show ?thesis by (simp add: A'_def add_ac) qed hence "size {# y \# r A'. y \ X #} \ size {# y \# r A. y \ X #}" by (simp add: committee_preference_def) finally show ?thesis by simp next assume l: "l > 1" define n where "n = count A X" have "n \ size A" unfolding n_def by (rule count_le_size) hence "n \ n_voters" by (simp add: A.size_A) define m where "m = nat (ceiling (n_voters / committee_size))" have "n_voters / committee_size \ m" unfolding m_def by linarith hence m: "n_voters \ committee_size * m" using \committee_size > 0\ by (simp add: field_simps flip: of_nat_mult) have "real n_voters / real committee_size > 0" using n_voters_pos less.prems by auto hence m': "\real n_voters / real committee_size\ = int m" by (simp add: m_def) have "1 * m \ l * m" using l by (intro mult_right_mono) auto also have "l * m \ n" using less.prems by (simp add: m' n_def flip: of_nat_mult) finally have "m \ n" by simp with less.prems l have "X \ {}" by auto then obtain x where x: "x \ X" by blast have "card (X - {x}) > 0" using less.prems x l by simp hence "X - {x} \ {}" by force have "count A X > 0" by (rule Nat.gr0I) (use n_voters_pos less.prems l in \auto simp: field_simps mult_le_0_iff\) hence "X \# A" by force have [simp]: "replicate_mset n X \# A" by (simp add: n_def flip: count_le_replicate_mset_subset_eq) define A'' where "A'' = A - replicate_mset n X" define A' where "A' = A'' + replicate_mset m {x} + replicate_mset (n - m) (X - {x})" interpret A': anon_papp_profile n_voters parties committee_size A' proof show "Y \ parties" if "Y \# A'" for Y using that A.A_subset x \X \# A\ by (fastforce simp: A'_def A''_def dest!: in_diffD split: if_splits) next show "{} \# A'" using A.A_nonempty x \X \# A\ \X - {x} \ {}\ by (auto simp: A'_def A''_def dest!: in_diffD split: if_splits) next show "size A' = n_voters" using \m \ n\ by (auto simp: A'_def A''_def A.size_A subset_mset.add_increasing2 size_Diff_submset) qed have "x \# r A'" proof (rule weak_representation) show "is_pref_profile A'" by (fact A'.anon_papp_profile_axioms) next have "n_voters \ committee_size * m" by (fact m) also have "m \ count A' {x}" by (simp add: A'_def) finally show "n_voters \ committee_size * count A' {x}" by simp qed hence "1 \ count (r A') x" by simp also have "\ = size {# y \# r A'. y = x #}" by simp finally have 1: "size {#y \# r A'. y = x#} \ 1" . have 2: "size {# y \# r A'. y \ X - {x} #} \ l - 1" proof (rule less.IH) have "int (l - 1) * \real n_voters / real committee_size\ = int ((l - 1) * m)" by (auto simp add: m' not_less) also have "(l - 1) * m = l * m - m" by (simp add: algebra_simps) also have "l * m \ n" using less.prems by (simp add: m' n_def flip: of_nat_mult) hence "l * m - m \ n - m" by (meson diff_le_mono) also have "n - m \ count A' (X - {x})" by (simp add: A'_def A''_def) finally show "int (l - 1) * \real n_voters / real committee_size\ \ int (count A' (X - {x}))" by simp qed (use l A'.anon_papp_profile_axioms x less.prems in \auto\) have "1 + (l - 1) \ size {#y \# r A'. y = x#} + size {#y \# r A'. y \ X - {x}#}" by (intro add_mono 1 2) also have "\ = size ({#y \# r A'. y = x#} + {#y \# r A'. y \ X - {x}#})" by simp also have "{#y \# r A'. y = x#} + {#y \# r A'. y \ X - {x}#} = {#y \# r A'. y = x \ y \ X - {x}#}" by (rule filter_mset_disjunction [symmetric]) auto also have "(\y. y = x \ y \ X - {x}) = (\y. y \ X)" using x by auto also have "1 + (l - 1) = l" using l by simp also have "r A' \[Comm(X)] r A" proof - have "r (replicate_mset m {x} + replicate_mset (n - m) (X - {x}) + A'') \[Comm(X)] r (replicate_mset n X + A'')" proof (rule proposition2) show " {} \# A'' + (replicate_mset m {x} + replicate_mset (n - m) (X - {x}))" using A'.A_nonempty by (auto simp: A'_def) show "\X'\# A'' + (replicate_mset m {x} + replicate_mset (n - m) (X - {x})). X' \ parties" using A'.A_subset x by (auto simp: A'_def dest: in_diffD) show "size A'' + n = n_voters" using \n \ n_voters\ by (auto simp: A''_def size_Diff_submset A.size_A) qed (use less.prems l \m \ n\ in auto) also have "replicate_mset n X + A'' = A" by (simp add: A''_def n_def flip: count_le_replicate_mset_subset_eq) finally show ?thesis by (simp add: A'_def add_ac) qed hence "size {# y \# r A'. y \ X #} \ size {# y \# r A. y \ X #}" by (simp add: committee_preference_def) finally show ?thesis by simp qed qed end subsection \Dividing the number of voters\ text \ If we have a PAPP rule that satisfies weak representation and cardinality strategyproofness, for $ln$ voters, we can turn it into one for \n\ voters. This is done by simply cloning each voter \l\ times. Consequently, if we have an impossibility result for $n$ voters, it also holds for any integer multiple of \n\. \ locale divide_voters_card_stratproof_weak_rep_anon_papp = card_stratproof_weak_rep_anon_papp "l * n_voters" parties committee_size r for l n_voters parties committee_size r begin definition lift_profile :: "'a set multiset \ 'a set multiset" where "lift_profile A = (\X\#A. replicate_mset l X)" sublocale lowered: anon_papp_election n_voters parties by standard (use n_voters_pos in auto) lemma l_pos: "l > 0" using n_voters_pos by auto lemma empty_in_lift_profile_iff [simp]: "{} \# lift_profile A \ {} \# A" using l_pos by (auto simp: lift_profile_def) lemma set_mset_lift_profile [simp]: "set_mset (lift_profile A) = set_mset A" using l_pos by (auto simp: lift_profile_def) lemma size_lift_profile: "size (lift_profile A) = l * size A" by (simp add: size_mset_sum_mset lift_profile_def image_mset.compositionality o_def) lemma count_lift_profile [simp]: "count (lift_profile A) x = l * count A x" unfolding lift_profile_def by (induction A) auto lemma is_pref_profile_lift_profile [intro]: assumes "lowered.is_pref_profile A" shows "is_pref_profile (lift_profile A)" proof - interpret anon_papp_profile n_voters parties committee_size A by fact show ?thesis using A_nonempty A_subset size_A by unfold_locales (auto simp: lift_profile_def size_mset_sum_mset image_mset.compositionality o_def) qed sublocale lowered: anon_papp n_voters parties committee_size "r \ lift_profile" proof fix A assume "lowered.is_pref_profile A" hence "is_pref_profile (lift_profile A)" by blast hence "is_committee (r (lift_profile A))" using rule_wf by blast thus "lowered.is_committee ((r \ lift_profile) A)" by simp qed sublocale lowered: weak_rep_anon_papp n_voters parties committee_size "r \ lift_profile" proof fix A x assume A: "lowered.is_pref_profile A" and x: "n_voters \ committee_size * count A {x}" from A have A': "is_pref_profile (lift_profile A)" by blast from x have "l * n_voters \ l * (committee_size * count A {x})" by (rule mult_left_mono) auto also have "\ = committee_size * count (lift_profile A) {x}" by simp finally have "x \# r (lift_profile A)" by (intro weak_representation A') thus "x \# (r \ lift_profile) A" by simp qed sublocale lowered: card_stratproof_anon_papp n_voters parties committee_size "r \ lift_profile" proof fix A X Y show "\lowered.card_manipulable A X Y" unfolding lowered.card_manipulable_def proof (rule notI, elim conjE) assume A: "lowered.is_pref_profile A" and XY: "X \# A" "Y \ {}" "Y \ parties" assume *: "(r \ lift_profile) A \[lowered.committee_preference X] (r \ lift_profile) (A - {#X#} + {#Y#})" interpret anon_papp_profile n_voters parties committee_size A by fact have X: "X \ {}" "X \ parties" using XY A_nonempty A_subset by auto define A' where "A' = A - {#X#}" have A': "A = A' + {#X#}" using XY by (simp add: A'_def) have "r (lift_profile A) \[committee_preference X] r (lift_profile (A - {#X#} + {#Y#}))" using * by simp also have "r (lift_profile (A - {#X#} + {#Y#})) \[committee_preference X] r (lift_profile A - {#X#} + {#Y#})" proof - have "r (replicate_mset (l - 1) Y + (lift_profile A' + {#Y#})) \[committee_preference X] r (replicate_mset (l - 1) X + (lift_profile A' + {#Y#}))" proof (rule proposition2) show "size (lift_profile A' + {#Y#}) + (l - 1) = l * n_voters" using XY l_pos n_voters_pos by (simp add: A'_def size_lift_profile size_Diff_singleton algebra_simps Suc_diff_le size_A) next show "{} \# lift_profile A' + {#Y#} + replicate_mset (l - 1) Y" using XY A_nonempty by (auto simp: A'_def dest: in_diffD) next show "\X'\#lift_profile A' + {#Y#} + replicate_mset (l - 1) Y. X' \ parties" using XY A_subset by (auto simp: A'_def dest: in_diffD) qed (use X in auto) thus ?thesis by (simp add: A' replicate_mset_rec l_pos lift_profile_def) qed finally have "card_manipulable (lift_profile A) X Y" unfolding card_manipulable_def using XY A by auto with not_manipulable show False by blast qed qed sublocale lowered: card_stratproof_weak_rep_anon_papp n_voters parties committee_size "r \ lift_profile" .. end locale divide_voters_card_stratproof_weak_prop_rep_anon_papp = card_stratproof_weak_prop_rep_anon_papp "l * n_voters" parties committee_size r for l n_voters parties committee_size r begin sublocale divide_voters_card_stratproof_weak_rep_anon_papp .. sublocale lowered: card_stratproof_weak_prop_rep_anon_papp n_voters parties committee_size "r \ lift_profile" proof fix A x l' assume A: "lowered.is_pref_profile A" and x: "l' * n_voters \ committee_size * count A {x}" from A have A': "is_pref_profile (lift_profile A)" by blast from x have "l * (l' * n_voters) \ l * (committee_size * count A {x})" by (rule mult_left_mono) auto also have "\ = committee_size * count (lift_profile A) {x}" by simp also have "l * (l' * n_voters) = l' * (l * n_voters)" by (simp add: algebra_simps) finally have "count (r (lift_profile A)) x \ l'" by (intro weak_proportional_representation A') thus "count ((r \ lift_profile) A) x \ l'" by simp qed end subsection \Decreasing the number of parties\ text \ If we have a PAPP rule that satisfies weak representation and cardinality strategyproofness, for $m$ parties, we can turn it into one for $m-1$ parties. This is done by simply duplicating one particular party (say \x\) in the preference profile, i.e.\ whenever \x\ is part of an approval list, we add a clone of \x\ (say \y\) as well. Should \y\ then end up in the committee, we simply replace it with \x\. Consequently, if we have an impossibility result for $k$ parties, it also holds for $\geq m$ parties. \ locale remove_alt_card_stratproof_weak_rep_anon_papp = card_stratproof_weak_rep_anon_papp n_voters parties committee_size r for n_voters and parties :: "'a set" and committee_size r + fixes x y :: 'a assumes xy: "x \ parties" "y \ parties" "x \ y" begin definition lift_applist :: "'a set \ 'a set" where "lift_applist X = (if x \ X then insert y X else X)" definition lift_profile :: "'a set multiset \ 'a set multiset" where "lift_profile A = image_mset lift_applist A" definition lower_result where "lower_result C = image_mset (\z. if z = y then x else z) C" definition lowered where "lowered = lower_result \ r \ lift_profile" lemma lift_profile_empty [simp]: "lift_profile {#} = {#}" by (simp add: lift_profile_def) lemma lift_profile_add_mset [simp]: "lift_profile (add_mset X A) = add_mset (lift_applist X) (lift_profile A)" by (simp add: lift_profile_def) lemma empty_in_lift_profile_iff [simp]: "{} \# lift_profile A \ {} \# A" by (auto simp: lift_applist_def lift_profile_def) lemma size_lift_profile [simp]: "size (lift_profile A) = size A" by (simp add: lift_profile_def) lemma lift_applist_eq_self_iff [simp]: "lift_applist X = X \ x \ X \ y \ X" by (auto simp: lift_applist_def) lemma lift_applist_eq_self_iff' [simp]: "lift_applist (X - {y}) = X \ (x \ X \ y \ X)" by (cases "y \ X") (auto simp: lift_applist_def xy) lemma in_lift_applist_iff: "z \ lift_applist X \ z \ X \ (z = y \ x \ X)" by (auto simp: lift_applist_def) lemma count_lift_profile: assumes "\Y\#A. y \ Y" shows "count (lift_profile A) X = (if x \ X \ y \ X then count A (X - {y}) else 0)" using assms xy by (induction A) (auto simp: lift_applist_def) lemma y_notin_lower_result [simp]: "y \# lower_result C" using xy by (auto simp: lower_result_def) lemma lower_result_subset: "set_mset (lower_result C) \ insert x (set_mset C - {y})" using xy by (auto simp: lower_result_def) lemma lower_result_subset': "set_mset C \ parties \ set_mset (lower_result C) \ parties" using xy by (auto simp: lower_result_def) lemma size_lower_result [simp]: "size (lower_result C) = size C" by (simp add: lower_result_def) lemma count_lower_result: "count (lower_result C) z = (if z = y then 0 else if z = x then count C x + count C y else count C z)" using xy by (induction C) (auto simp: lower_result_def) lemma in_lower_result_iff: "z \# lower_result C \ z \ y \ (z \# C \ (z = x \ y \# C))" unfolding lower_result_def using xy by (induction C) auto sublocale lowered: anon_papp_election n_voters "parties - {y}" by standard (use n_voters_pos xy in auto) lemma is_pref_profile_lift_profile [intro]: assumes "lowered.is_pref_profile A" shows "is_pref_profile (lift_profile A)" proof - interpret anon_papp_profile n_voters "parties - {y}" committee_size A by fact show ?thesis using A_nonempty A_subset size_A by unfold_locales (auto simp: lift_profile_def lift_applist_def xy size_mset_sum_mset image_mset.compositionality o_def) qed sublocale lowered: anon_papp n_voters "parties - {y}" committee_size lowered proof fix A assume "lowered.is_pref_profile A" hence "is_pref_profile (lift_profile A)" by blast hence "is_committee (r (lift_profile A))" using rule_wf by blast thus "lowered.is_committee (lowered A)" unfolding lowered.is_committee_def is_committee_def lowered_def using lower_result_subset'[of "r (lift_profile A)"] by auto qed sublocale lowered: weak_rep_anon_papp n_voters "parties - {y}" committee_size lowered proof fix A z assume A: "lowered.is_pref_profile A" and z: "n_voters \ committee_size * count A {z}" interpret A: anon_papp_profile n_voters "parties - {y}" committee_size A by fact have "committee_size > 0" using z n_voters_pos by (intro Nat.gr0I) auto from A have A': "is_pref_profile (lift_profile A)" by blast have "count A {z} > 0" using z n_voters_pos by (intro Nat.gr0I) auto hence "{z} \# A" by simp hence z': "z \ parties - {y}" using A.A_subset z by auto define C where "C = r (lift_profile A)" show "z \# lowered A" proof (cases "z = x") case False have "n_voters \ committee_size * count A {z}" by fact also have "count A {z} \ count (lift_profile A) {z}" using A.A_subset z' False by (subst count_lift_profile) auto hence "committee_size * count A {z} \ committee_size * count (lift_profile A) {z}" by (intro mult_left_mono) auto finally have "z \# r (lift_profile A)" by (intro weak_representation A') thus "z \# lowered A" using False z' by (simp add: lowered_def in_lower_result_iff) next case [simp]: True have "1 \ size {#z \# C. z \ {x, y}#}" unfolding C_def proof (rule proposition3) have [simp]: "{x, y} - {y} = {x}" using xy by auto hence "n_voters \ committee_size * count (lift_profile A) {x, y}" using xy A.A_subset z by (subst count_lift_profile) auto thus "int 1 * \real n_voters / real committee_size\ \ int (count (lift_profile A) {x, y})" using \committee_size > 0\ by (auto simp: ceiling_le_iff field_simps simp flip: of_nat_mult) qed (use A' xy \committee_size > 0\ in auto) also have "\ = count C x + count C y" using xy by (induction C) auto also have "\ = count (lowered A) x" using xy by (simp add: lowered_def count_lower_result C_def) finally show "z \# lowered A" by simp qed qed lemma filter_lower_result_eq: "y \ X \ {#x \# lower_result C. x \ X#} = lower_result {#x \# C. x \ lift_applist X#}" by (induction C) (auto simp: lower_result_def lift_applist_def) sublocale lowered: card_stratproof_anon_papp n_voters "parties - {y}" committee_size lowered proof fix A X Y show "\lowered.card_manipulable A X Y" unfolding lowered.card_manipulable_def proof (rule notI, elim conjE) assume A: "lowered.is_pref_profile A" and XY: "X \# A" "Y \ {}" "Y \ parties - {y}" assume *: "lowered A \[lowered.committee_preference X] lowered (A - {#X#} + {#Y#})" interpret anon_papp_profile n_voters "parties - {y}" committee_size A by fact have X: "X \ {}" "X \ parties - {y}" using XY A_nonempty A_subset by auto define A' where "A' = A - {#X#}" have A': "A = A' + {#X#}" using XY by (simp add: A'_def) from * have "size {#x \# lower_result (r (lift_profile A' + {#lift_applist X#})). x \ X#} < size {#x \# lower_result (r (lift_profile A' + {#lift_applist Y#})). x \ X#}" by (simp add: lowered_def A' lowered.strong_committee_preference_iff) also have "{#x \# lower_result (r (lift_profile A' + {#lift_applist X#})). x \ X#} = lower_result {#x \# r (lift_profile A' + {#lift_applist X#}). x \ lift_applist X#}" using X by (subst filter_lower_result_eq) auto also have "{#x \# lower_result (r (lift_profile A' + {#lift_applist Y#})). x \ X#} = lower_result {#x \# r (lift_profile A' + {#lift_applist Y#}). x \ lift_applist X#}" using X by (subst filter_lower_result_eq) auto finally have "size {#x \# r (lift_profile A' + {#lift_applist X#}). x \ lift_applist X#} < size {#x \# r (lift_profile A' + {#lift_applist Y#}). x \ lift_applist X#}" by simp hence "r (lift_profile A' + {#lift_applist X#}) \[committee_preference (lift_applist X)] r (lift_profile A' + {#lift_applist Y#})" by (simp add: strong_committee_preference_iff) moreover have "\r (lift_profile A' + {#lift_applist X#}) \[committee_preference (lift_applist X)] r (lift_profile A' + {#lift_applist Y#})" proof (rule not_manipulable' [where Y = "lift_applist Y"]) have "is_pref_profile (lift_profile A)" using A by blast thus "is_pref_profile (lift_profile A' + {#lift_applist X#})" using A by (simp add: A') next have "is_pref_profile (lift_profile (A - {#X#} + {#Y#}))" using A XY lowered.is_pref_profile_replace by blast thus "is_pref_profile (lift_profile A' + {#lift_applist Y#})" by (simp add: A') qed auto ultimately show False by contradiction qed qed sublocale lowered: card_stratproof_weak_rep_anon_papp n_voters "parties - {y}" committee_size lowered .. end -(* -locale remove_alt_card_stratproof_weak_prop_rep_anon_papp = - card_stratproof_weak_prop_rep_anon_papp n_voters parties committee_size r - for n_voters and parties :: "'a set" and committee_size r + - fixes x y :: 'a - assumes xy: "x \ parties" "y \ parties" "x \ y" -begin - -sublocale remove_alt_card_stratproof_weak_rep_anon_papp - by standard (use xy in auto) - -sublocale lowered: weak_prop_rep_anon_papp n_voters "parties - {y}" committee_size lowered -proof - fix A z l - assume A: "lowered.is_pref_profile A" and z: "l * n_voters \ committee_size * count A {z}" - interpret A: anon_papp_profile n_voters "parties - {y}" committee_size A - by fact - - show "count (lowered A) z \ l" - proof (cases "l = 0") - case True - thus ?thesis - by simp - next - case l: False - - have "committee_size > 0" - using z n_voters_pos l by (intro Nat.gr0I) auto - - from A have A': "is_pref_profile (lift_profile A)" - by blast - have "count A {z} > 0" - using z n_voters_pos l by (intro Nat.gr0I) auto - hence "{z} \# A" - by simp - hence z': "z \ parties - {y}" - using A.A_subset z by auto - - define C where "C = r (lift_profile A)" - - show ?thesis - proof (cases "z = x") - case False - have "l * n_voters \ committee_size * count A {z}" - by fact - also have "count A {z} \ count (lift_profile A) {z}" - using A.A_subset z' False by (subst count_lift_profile) auto - hence "committee_size * count A {z} \ committee_size * count (lift_profile A) {z}" - by (intro mult_left_mono) auto - finally have "count (r (lift_profile A)) z \ l" - by (intro weak_proportional_representation A') - thus ?thesis - using False z' by (simp add: lowered_def count_lower_result) - next - case [simp]: True - show ?thesis - sorry - qed - qed -qed - -end -*) - - text \ The following lemma is now simply an iterated application of the above. This allows us to restrict a P-APP rule to any non-empty subset of parties. \ lemma card_stratproof_weak_rep_anon_papp_restrict_parties: assumes "card_stratproof_weak_rep_anon_papp n parties k r" "parties' \ parties" "parties' \ {}" shows "\r. card_stratproof_weak_rep_anon_papp n parties' k r" proof - have "finite parties" proof - interpret card_stratproof_weak_rep_anon_papp n parties k r by fact show ?thesis by (rule finite_parties) qed thus ?thesis using assms proof (induction parties arbitrary: r rule: finite_psubset_induct) case (psubset parties r) show ?thesis proof (cases "parties = parties'") case True thus ?thesis using psubset.prems by auto next case False obtain x where x: "x \ parties'" using psubset.prems by blast from False and psubset.prems have "parties - parties' \ {}" by auto then obtain y where y: "y \ parties - parties'" by blast interpret card_stratproof_weak_rep_anon_papp n parties k r by fact interpret remove_alt_card_stratproof_weak_rep_anon_papp n parties k r x y by standard (use x y psubset.prems in auto) show ?thesis proof (rule psubset.IH) show "parties - {y} \ parties" and "parties' \ parties - {y}" "parties' \ {}" using x y psubset.prems by auto next show "card_stratproof_weak_rep_anon_papp n (parties - {y}) k lowered" using lowered.card_stratproof_weak_rep_anon_papp_axioms . qed qed qed qed subsection \Decreasing the committee size\ text \ If we have a PAPP rule that satisfies weak representation and cardinality strategyproofness, for $l(k+1)$ voters, $m+1$ parties, and a committee size of $k+1$, we can turn it into one for $lk$ voters, $m$ parties, and a committee size of $k$. This is done by again cloning a party \x\ into a new party \y\ and additionally adding \l\ new voters whose preferences are \{x, y}\. We again replace any \y\ occuring in the output committee with \x\. Weak representation then ensures that \x\ occurs in the output at least once, and we then simply remove one \x\ from the committee to obtain an output committee of size \k - 1\. Consequently, if we have an impossibility result for a committee size of $m$, we can extend it to a larger committee size, but at the cost of introducing a new party and new voters, and with a restriction on the number of voters. \ (* n / k = l *) locale decrease_committee_card_stratproof_weak_rep_anon_papp = card_stratproof_weak_rep_anon_papp "l * (k + 1)" "insert y parties" "k + 1" r for l k y and parties :: "'a set" and r + fixes x :: 'a assumes xy: "x \ parties" "y \ parties" assumes k: "k > 0" begin definition lift_applist :: "'a set \ 'a set" where "lift_applist X = (if x \ X then insert y X else X)" definition lift_profile :: "'a set multiset \ 'a set multiset" where "lift_profile A = image_mset lift_applist A + replicate_mset l {x, y}" definition lower_result where "lower_result C = image_mset (\z. if z = y then x else z) C - {#x#}" definition lowered where "lowered = lower_result \ r \ lift_profile" lemma l: "l > 0" using n_voters_pos by auto lemma x_neq_y [simp]: "x \ y" "y \ x" using xy by auto lemma lift_profile_empty [simp]: "lift_profile {#} = replicate_mset l {x, y}" by (simp add: lift_profile_def) lemma lift_profile_add_mset [simp]: "lift_profile (add_mset X A) = add_mset (lift_applist X) (lift_profile A)" by (simp add: lift_profile_def) lemma empty_in_lift_profile_iff [simp]: "{} \# lift_profile A \ {} \# A" by (auto simp: lift_applist_def lift_profile_def) lemma size_lift_profile [simp]: "size (lift_profile A) = size A + l" by (simp add: lift_profile_def) lemma lift_applist_eq_self_iff [simp]: "lift_applist X = X \ x \ X \ y \ X" by (auto simp: lift_applist_def) lemma lift_applist_eq_self_iff' [simp]: "lift_applist (X - {y}) = X \ (x \ X \ y \ X)" by (cases "y \ X") (auto simp: lift_applist_def xy) lemma in_lift_applist_iff: "z \ lift_applist X \ z \ X \ (z = y \ x \ X)" by (auto simp: lift_applist_def) lemma count_lift_profile: assumes "\Y\#A. y \ Y" shows "count (lift_profile A) X = (if x \ X \ y \ X then count A (X - {y}) else 0) + (if X = {x, y} then l else 0)" using assms xy by (induction A) (auto simp: lift_applist_def) lemma y_notin_lower_result [simp]: "y \# lower_result C" using xy by (auto simp: lower_result_def dest: in_diffD) lemma lower_result_subset: "set_mset (lower_result C) \ insert x (set_mset C - {y})" using xy by (auto simp: lower_result_def dest: in_diffD) lemma lower_result_subset': "set_mset C \ insert y parties \ set_mset (lower_result C) \ parties" by (rule order.trans[OF lower_result_subset]) (use xy in auto) lemma size_lower_result [simp]: assumes "x \# C \ y \# C" shows "size (lower_result C) = size C - 1" using assms by (auto simp: lower_result_def size_Diff_singleton) lemma size_lower_result': "size (lower_result C) = size C - (if x \# C \ y \# C then 1 else 0)" proof - define f where "f = (\C. image_mset (\z. if z = y then x else z) C)" have "size (lower_result C) = size (f C - {#x#})" by (simp add: lower_result_def f_def) also have "\ = size (f C) - (if x \# f C then 1 else 0)" by (simp add: diff_single_trivial size_Diff_singleton) finally show ?thesis by (auto simp: f_def) qed lemma count_lower_result: "count (lower_result C) z = (if z = y then 0 else if z = x then count C x + count C y - 1 else count C z)" (is "_ = ?rhs") proof - define f where "f = (\C. image_mset (\z. if z = y then x else z) C)" have "count (lower_result C) z = count (f C - {#x#}) z" by (simp add: lower_result_def f_def) also have "\ = count (f C) z - (if z = x then 1 else 0)" by auto also have "count (f C) z = (if z = y then 0 else if z = x then count C x + count C y else count C z)" using xy by (induction C) (auto simp: f_def) also have "\ - (if z = x then 1 else 0) = ?rhs" by auto finally show ?thesis . qed lemma in_lower_resultD: "z \# lower_result C \ z = x \ z \# C" using xy by (auto simp: lower_result_def dest!: in_diffD) lemma in_lower_result_iff: "z \# lower_result C \ z \ y \ (if z = x then count C x + count C y > 1 else z \# C)" (is "_ = ?rhs") proof - have "z \# lower_result C \ count (lower_result C) z > 0" by auto also have "\ \ ?rhs" by (subst count_lower_result) auto finally show ?thesis . qed lemma filter_lower_result_eq: assumes "y \ X" shows "{#z \# lower_result C. z \ X#} = lower_result {#z \# C. z \ lift_applist X#}" proof - define f where "f = (\C. {#if z = y then x else z. z \# C#})" have "lower_result {#z \# C. z \ lift_applist X#} = f {#z \# C. z \ lift_applist X#} - {#x#}" by (simp add: f_def lower_result_def) also have "f {#z \# C. z \ lift_applist X#} = {#z \# f C. z \ X#}" using assms by (induction C) (auto simp: f_def lift_applist_def) also have "\ - {#x#} = {#z \# f C - {#x#}. z \ X#}" by (subst filter_diff_mset') auto also have "f C - {#x#} = lower_result C" by (simp add: f_def lower_result_def) finally show ?thesis .. qed sublocale lowered: anon_papp_election "l * k" parties k by standard (use n_voters_pos xy finite_parties k in auto) lemma is_pref_profile_lift_profile [intro]: assumes "lowered.is_pref_profile A" shows "is_pref_profile (lift_profile A)" proof - interpret A: anon_papp_profile "l * k" parties k A by fact show ?thesis using A.A_nonempty A.A_subset A.size_A l by unfold_locales (auto simp: lift_profile_def lift_applist_def xy size_mset_sum_mset image_mset.compositionality o_def) qed lemma is_committee_lower_result: assumes "is_committee C" "x \# C \ y \# C" shows "lowered.is_committee (lower_result C)" using assms unfolding is_committee_def lowered.is_committee_def using lower_result_subset'[of C] by auto lemma x_or_y_in_r_lift_profile: assumes "lowered.is_pref_profile A" shows "x \# r (lift_profile A) \ y \# r (lift_profile A)" proof - interpret A: anon_papp_profile "l * k" parties k A by fact have "size {#z \# r (lift_profile A). z \ {x, y}#} \ 1" proof (rule proposition3) have "real (l * (k + 1)) / real (k + 1) = real l" by (simp add: field_simps) also have "int 1 * \\\ = int l" by simp also have "l \ count (lift_profile A) {x, y}" using xy A.A_subset by (subst count_lift_profile) auto finally show "int 1 * \real (l * (k + 1)) / real (k + 1)\ \ int (count (lift_profile A) {x, y})" by simp next show "is_pref_profile (lift_profile A)" by (intro is_pref_profile_lift_profile) fact qed (use xy in auto) hence "{#z \# r (lift_profile A). z \ {x, y}#} \ {#}" by force thus ?thesis by auto qed sublocale lowered: anon_papp "l * k" parties k lowered proof fix A assume A: "lowered.is_pref_profile A" hence "is_pref_profile (lift_profile A)" by blast hence "is_committee (r (lift_profile A))" using rule_wf by blast thus "lowered.is_committee (lowered A)" unfolding lowered_def o_def using x_or_y_in_r_lift_profile[of A] A by (intro is_committee_lower_result) auto qed sublocale lowered: weak_rep_anon_papp "l * k" parties k lowered proof fix A z assume A: "lowered.is_pref_profile A" and z: "l * k \ k * count A {z}" interpret A: anon_papp_profile "l * k" "parties" k A by fact from A have A': "is_pref_profile (lift_profile A)" by blast have "count A {z} > 0" using z k n_voters_pos by (intro Nat.gr0I) auto hence "{z} \# A" by simp hence z': "z \ parties" using A.A_subset z by auto hence [simp]: "y \ z" "z \ y" using xy by auto define C where "C = r (lift_profile A)" show "z \# lowered A" proof (cases "z = x") case False have "l \ count A {z}" using z k by (simp add: algebra_simps) hence "l * (k + 1) \ (k + 1) * count A {z}" by (subst mult.commute, intro mult_right_mono) auto also have "count A {z} = count (lift_profile A) {z}" using False A.A_subset xy by (subst count_lift_profile) auto finally have "z \# r (lift_profile A)" by (intro weak_representation A') thus "z \# lowered A" using False by (auto simp: lowered_def in_lower_result_iff) next case [simp]: True from xy have [simp]: "{x, y} - {y} = {x}" by auto have "size {#z \# C. z \ {x, y}#} \ 2" unfolding C_def proof (rule proposition3) have "real (l * (k + 1)) / real (k + 1) = l" unfolding of_nat_mult using k by (simp add: divide_simps) also have "int 2 * \\\ = int (2 * l)" by simp also have "\ \ count (lift_profile A) {x, y}" using z k xy A.A_subset by (subst count_lift_profile) auto finally show "int 2 * \real (l * (k + 1)) / real (k + 1)\ \ \" . qed (use A' xy in auto) also have "size {#z \# C. z \ {x, y}#} = count C x + count C y" by (induction C) auto finally have "x \# lower_result C" by (subst in_lower_result_iff) auto thus "z \# lowered A" by (simp add: lowered_def C_def) qed qed sublocale lowered: card_stratproof_anon_papp "l * k" parties k lowered proof fix A X Y show "\lowered.card_manipulable A X Y" unfolding lowered.card_manipulable_def proof (rule notI, elim conjE) assume A: "lowered.is_pref_profile A" and XY: "X \# A" "Y \ {}" "Y \ parties" assume *: "lowered A \[lowered.committee_preference X] lowered (A - {#X#} + {#Y#})" interpret anon_papp_profile "l * k" "parties" k A by fact have X: "X \ {}" "X \ parties" using XY A_nonempty A_subset by auto define A' where "A' = A - {#X#}" have A': "A = A' + {#X#}" using XY by (simp add: A'_def) from xy X XY have [simp]: "y \ X" "y \ Y" by auto define Al1 where "Al1 = lift_profile A" define Al2 where "Al2 = lift_profile (A' + {#Y#})" have A'_plus_Y: "lowered.is_pref_profile (A' + {#Y#})" unfolding A'_def using A XY lowered.is_pref_profile_replace by blast have Al1: "is_pref_profile Al1" unfolding Al1_def using A by blast have Al2: "is_pref_profile Al2" unfolding Al2_def unfolding A'_def using A XY lowered.is_pref_profile_replace by blast have size_aux: "size (lower_result {#x \# r (lift_profile A). x \ lift_applist X#}) = size {#x \# r (lift_profile A). x \ lift_applist X#} - (if x \ X then 1 else 0)" if A: "lowered.is_pref_profile A" for A using x_or_y_in_r_lift_profile[OF A] by (subst size_lower_result') (auto simp: lift_applist_def) from * have "size {#x \# lower_result (r Al1). x \ X#} < size {#x \# lower_result (r Al2). x \ X#}" by (simp add: Al1_def Al2_def lowered_def A' lowered.strong_committee_preference_iff) also have "{#x \# lower_result (r Al1). x \ X#} = lower_result {#x \# r Al1. x \ lift_applist X#}" using X xy by (subst filter_lower_result_eq) auto also have "{#x \# lower_result (r Al2). x \ X#} = lower_result {#x \# r Al2. x \ lift_applist X#}" using X xy by (subst filter_lower_result_eq) auto also have "size (lower_result {#x \# r Al1. x \ lift_applist X#}) = size {#x \# r Al1. x \ lift_applist X#} - (if x \ X then 1 else 0)" unfolding Al1_def by (rule size_aux) fact also have "size (lower_result {#x \# r Al2. x \ lift_applist X#}) = size {#x \# r Al2. x \ lift_applist X#} - (if x \ X then 1 else 0)" unfolding Al2_def by (rule size_aux) fact finally have "size {#x \# r Al1. x \ lift_applist X#} < size {#x \# r Al2. x \ lift_applist X#}" by auto hence "r Al1 \[committee_preference (lift_applist X)] r Al2" by (simp add: strong_committee_preference_iff) moreover have "\r Al1 \[committee_preference (lift_applist X)] r Al2" by (rule not_manipulable' [where Y = "lift_applist Y"]) (use Al1 Al2 in \auto simp: Al1_def Al2_def A'\) ultimately show False by contradiction qed qed sublocale lowered: card_stratproof_weak_rep_anon_papp "l * k" parties k lowered .. end text \ For Weak \<^emph>\Proportional\ Representation, the lowering argument to decrease the committee size is somewhat easier since it does not involve adding a new party; instead, we simply add \l\ new voters whose preferences are \{x}\. \ locale decrease_committee_card_stratproof_weak_prop_rep_anon_papp = card_stratproof_weak_prop_rep_anon_papp "l * (k + 1)" parties "k + 1" r for l k and parties :: "'a set" and r + fixes x :: 'a assumes x: "x \ parties" assumes k: "k > 0" begin definition lift_profile :: "'a set multiset \ 'a set multiset" where "lift_profile A = A + replicate_mset l {x}" definition lower_result where "lower_result C = C - {#x#}" definition lowered where "lowered = lower_result \ r \ lift_profile" lemma l: "l > 0" using n_voters_pos by auto lemma lift_profile_empty [simp]: "lift_profile {#} = replicate_mset l {x}" by (simp add: lift_profile_def) lemma lift_profile_add_mset [simp]: "lift_profile (add_mset X A) = add_mset X (lift_profile A)" by (simp add: lift_profile_def) lemma empty_in_lift_profile_iff [simp]: "{} \# lift_profile A \ {} \# A" by (auto simp: lift_profile_def) lemma size_lift_profile [simp]: "size (lift_profile A) = size A + l" by (simp add: lift_profile_def) lemma count_lift_profile: "count (lift_profile A) X = count A X + (if X = {x} then l else 0)" by (auto simp: lift_profile_def) lemma size_lower_result [simp]: assumes "x \# C" shows "size (lower_result C) = size C - 1" using assms by (auto simp: lower_result_def size_Diff_singleton) lemma size_lower_result': "size (lower_result C) = size C - (if x \# C then 1 else 0)" by (induction C) (auto simp: lower_result_def size_Diff_singleton) lemma count_lower_result: "count (lower_result C) z = count C z - (if z = x then 1 else 0)" by (auto simp: lower_result_def) lemma in_lower_resultD: "z \# lower_result C \ z \# C" by (auto simp: lower_result_def dest!: in_diffD) lemma in_lower_result_iff: "z \# lower_result C \ (if z = x then count C x > 1 else z \# C)" (is "_ = ?rhs") proof - have "z \# lower_result C \ count (lower_result C) z > 0" by auto also have "\ \ ?rhs" by (subst count_lower_result) auto finally show ?thesis . qed sublocale lowered: anon_papp_election "l * k" parties k by standard (use n_voters_pos finite_parties k in auto) lemma is_pref_profile_lift_profile [intro]: assumes "lowered.is_pref_profile A" shows "is_pref_profile (lift_profile A)" proof - interpret A: anon_papp_profile "l * k" parties k A by fact show ?thesis using A.A_nonempty A.A_subset A.size_A l by unfold_locales (auto simp: lift_profile_def x size_mset_sum_mset image_mset.compositionality o_def) qed lemma is_committee_lower_result: assumes "is_committee C" "x \# C" shows "lowered.is_committee (lower_result C)" using assms unfolding is_committee_def lowered.is_committee_def by (auto simp: lower_result_def size_Diff_singleton dest: in_diffD) lemma x_in_r_lift_profile: assumes "lowered.is_pref_profile A" shows "x \# r (lift_profile A)" proof (rule weak_representation) show "is_pref_profile (lift_profile A)" using assms by blast next have "(k + 1) * l \ (k + 1) * count (lift_profile A) {x}" by (intro mult_left_mono) (auto simp: count_lift_profile) thus "l * (k + 1) \ (k + 1) * count (lift_profile A) {x}" by (simp add: algebra_simps) qed sublocale lowered: anon_papp "l * k" parties k lowered proof fix A assume A: "lowered.is_pref_profile A" hence "is_pref_profile (lift_profile A)" by blast hence "is_committee (r (lift_profile A))" using rule_wf by blast thus "lowered.is_committee (lowered A)" unfolding lowered_def o_def using x_in_r_lift_profile[of A] A by (intro is_committee_lower_result) auto qed sublocale lowered: weak_prop_rep_anon_papp "l * k" parties k lowered proof fix A z l' assume A: "lowered.is_pref_profile A" and z: "l' * (l * k) \ k * count A {z}" interpret A: anon_papp_profile "l * k" "parties" k A by fact show "count (lowered A) z \ l'" proof (cases "l' > 0") case False thus ?thesis by auto next case l: True from A have A': "is_pref_profile (lift_profile A)" by blast have "count A {z} > 0" using z k n_voters_pos l by (intro Nat.gr0I) auto hence "{z} \# A" by simp hence z': "z \ parties" using A.A_subset z by auto define C where "C = r (lift_profile A)" show "count (lowered A) z \ l'" proof (cases "z = x") case False have "l' * l \ count A {z}" using z k by (simp add: algebra_simps) hence "l' * l * (k + 1) \ (k + 1) * count A {z}" by (subst mult.commute, intro mult_right_mono) auto also have "count A {z} = count (lift_profile A) {z}" using False A.A_subset by (subst count_lift_profile) auto finally have "count (r (lift_profile A)) z \ l'" by (intro weak_proportional_representation A') (auto simp: algebra_simps) thus "l' \ count (lowered A) z" using False by (simp add: lowered_def lower_result_def) next case [simp]: True have "l' * l \ count A {x}" using z k by (simp add: algebra_simps) hence "l' * l * (k + 1) \ (k + 1) * count A {x}" by (subst mult.commute, intro mult_right_mono) auto also have "\ + (k + 1) * l = (k + 1) * count (lift_profile A) {x}" by (simp add: count_lift_profile algebra_simps) finally have "(l' + 1) * (l * (k + 1)) \ (k + 1) * count (lift_profile A) {x}" by (simp add: algebra_simps) hence "count (r (lift_profile A)) x \ l' + 1" by (intro weak_proportional_representation A') thus "l' \ count (lowered A) z" by (simp add: lowered_def lower_result_def) qed qed qed sublocale lowered: card_stratproof_anon_papp "l * k" parties k lowered proof fix A X Y show "\lowered.card_manipulable A X Y" unfolding lowered.card_manipulable_def proof (rule notI, elim conjE) assume A: "lowered.is_pref_profile A" and XY: "X \# A" "Y \ {}" "Y \ parties" assume *: "lowered A \[lowered.committee_preference X] lowered (A - {#X#} + {#Y#})" interpret anon_papp_profile "l * k" "parties" k A by fact have X: "X \ {}" "X \ parties" using XY A_nonempty A_subset by auto define A' where "A' = A - {#X#}" have A': "A = A' + {#X#}" using XY by (simp add: A'_def) define Al1 where "Al1 = lift_profile A" define Al2 where "Al2 = lift_profile (A' + {#Y#})" have A'_plus_Y: "lowered.is_pref_profile (A' + {#Y#})" unfolding A'_def using A XY lowered.is_pref_profile_replace by blast have Al1: "is_pref_profile Al1" unfolding Al1_def using A by blast have Al2: "is_pref_profile Al2" unfolding Al2_def unfolding A'_def using A XY lowered.is_pref_profile_replace by blast have size_aux: "size (lower_result {#x \# r (lift_profile A). x \ X#}) = size {#x \# r (lift_profile A). x \ X#} - (if x \ X then 1 else 0)" if A: "lowered.is_pref_profile A" for A using x_in_r_lift_profile[OF A] by (subst size_lower_result') auto from * have "size {#x \# lower_result (r Al1). x \ X#} < size {#x \# lower_result (r Al2). x \ X#}" by (simp add: Al1_def Al2_def lowered_def A' lowered.strong_committee_preference_iff) also have "{#x \# lower_result (r Al1). x \ X#} = lower_result {#x \# r Al1. x \ X#}" using X x unfolding lower_result_def by (subst filter_diff_mset') auto also have "{#x \# lower_result (r Al2). x \ X#} = lower_result {#x \# r Al2. x \ X#}" using X x unfolding lower_result_def by (subst filter_diff_mset') auto also have "size (lower_result {#x \# r Al1. x \ X#}) = size {#x \# r Al1. x \ X#} - (if x \ X then 1 else 0)" unfolding Al1_def by (rule size_aux) fact also have "size (lower_result {#x \# r Al2. x \ X#}) = size {#x \# r Al2. x \ X#} - (if x \ X then 1 else 0)" unfolding Al2_def by (rule size_aux) fact finally have "size {#x \# r Al1. x \ X#} < size {#x \# r Al2. x \ X#}" by auto hence "r Al1 \[committee_preference X] r Al2" by (simp add: strong_committee_preference_iff) moreover have "\r Al1 \[committee_preference X] r Al2" by (rule not_manipulable' [where Y = "Y"]) (use Al1 Al2 in \auto simp: Al1_def Al2_def A'\) ultimately show False by contradiction qed qed sublocale lowered: card_stratproof_weak_prop_rep_anon_papp "l * k" parties k lowered .. end end diff --git a/thys/PAPP_Impossibility/ROOT b/thys/PAPP_Impossibility/ROOT --- a/thys/PAPP_Impossibility/ROOT +++ b/thys/PAPP_Impossibility/ROOT @@ -1,15 +1,14 @@ chapter AFP -session PAPP_Impossibility (AFP) = "HOL" + +session PAPP_Impossibility (AFP) = "HOL-Library" + options [timeout = 900] sessions - "HOL-Library" Randomised_Social_Choice theories [document=false] SAT_Replay theories PAPP_Impossibility document_files root.tex root.bib diff --git a/web/index.xml b/web/index.xml --- a/web/index.xml +++ b/web/index.xml @@ -1,13112 +1,13112 @@ Archive of Formal Proofs / Recent content on Archive of Formal Proofs Hugo -- gohugo.io en-gb Fri, 11 Nov 2022 00:00:00 +0000 Combinatorial Enumeration Algorithms /entries/Combinatorial_Enumeration_Algorithms.html Fri, 11 Nov 2022 00:00:00 +0000 /entries/Combinatorial_Enumeration_Algorithms.html The Balog–Szemerédi–Gowers Theorem /entries/Balog_Szemeredi_Gowers.html Fri, 11 Nov 2022 00:00:00 +0000 /entries/Balog_Szemeredi_Gowers.html The Incompatibility of Strategy-Proofness and Representation in Party-Approval Multi-Winner Elections /entries/PAPP_Impossibility.html Thu, 10 Nov 2022 00:00:00 +0000 /entries/PAPP_Impossibility.html Verification of Query Optimization Algorithms /entries/Query_Optimization.html Tue, 04 Oct 2022 00:00:00 +0000 /entries/Query_Optimization.html Maximum Segment Sum /entries/Maximum_Segment_Sum.html Thu, 29 Sep 2022 00:00:00 +0000 /entries/Maximum_Segment_Sum.html Undirected Graph Theory /entries/Undirected_Graph_Theory.html Thu, 29 Sep 2022 00:00:00 +0000 /entries/Undirected_Graph_Theory.html Making Arbitrary Relational Calculus Queries Safe-Range /entries/Safe_Range_RC.html Wed, 28 Sep 2022 00:00:00 +0000 /entries/Safe_Range_RC.html Stalnaker's Epistemic Logic /entries/Stalnaker_Logic.html Fri, 23 Sep 2022 00:00:00 +0000 /entries/Stalnaker_Logic.html p-adic Fields and p-adic Semialgebraic Sets /entries/Padic_Field.html Thu, 22 Sep 2022 00:00:00 +0000 /entries/Padic_Field.html Risk-Free Lending /entries/Risk_Free_Lending.html Sun, 18 Sep 2022 00:00:00 +0000 /entries/Risk_Free_Lending.html Soundness and Completeness of Implicational Logic /entries/Implicational_Logic.html Tue, 13 Sep 2022 00:00:00 +0000 /entries/Implicational_Logic.html CRYSTALS-Kyber /entries/CRYSTALS-Kyber.html Thu, 08 Sep 2022 00:00:00 +0000 /entries/CRYSTALS-Kyber.html Unbounded Separation Logic /entries/Separation_Logic_Unbounded.html Mon, 05 Sep 2022 00:00:00 +0000 /entries/Separation_Logic_Unbounded.html Khovanskii&#x27;s Theorem /entries/Khovanskii_Theorem.html Fri, 02 Sep 2022 00:00:00 +0000 /entries/Khovanskii_Theorem.html The Hales–Jewett Theorem /entries/Hales_Jewett.html Fri, 02 Sep 2022 00:00:00 +0000 /entries/Hales_Jewett.html Number Theoretic Transform /entries/Number_Theoretic_Transform.html Thu, 18 Aug 2022 00:00:00 +0000 /entries/Number_Theoretic_Transform.html Correctness of a Set-based Algorithm for Computing Strongly Connected Components of a Graph /entries/SCC_Bloemen_Sequential.html Wed, 17 Aug 2022 00:00:00 +0000 /entries/SCC_Bloemen_Sequential.html From THE BOOK: Two Squares via Involutions /entries/Involutions2Squares.html Mon, 15 Aug 2022 00:00:00 +0000 /entries/Involutions2Squares.html Verified Complete Test Strategies for Finite State Machines /entries/FSM_Tests.html Tue, 09 Aug 2022 00:00:00 +0000 /entries/FSM_Tests.html Nano JSON: Working with JSON formatted data in Isabelle/HOL and Isabelle/ML /entries/Nano_JSON.html Fri, 29 Jul 2022 00:00:00 +0000 /entries/Nano_JSON.html Isabelle/Solidity: A deep Embedding of Solidity in Isabelle/HOL /entries/Solidity.html Mon, 18 Jul 2022 00:00:00 +0000 /entries/Solidity.html Simultaneous diagonalization of pairwise commuting Hermitian matrices /entries/Commuting_Hermitian.html Mon, 18 Jul 2022 00:00:00 +0000 /entries/Commuting_Hermitian.html Pólya’s Proof of the Weighted Arithmetic–Geometric Mean Inequality /entries/Weighted_Arithmetic_Geometric_Mean.html Mon, 11 Jul 2022 00:00:00 +0000 /entries/Weighted_Arithmetic_Geometric_Mean.html A Reuse-Based Multi-Stage Compiler Verification for Language IMP /entries/IMP_Compiler_Reuse.html Sun, 10 Jul 2022 00:00:00 +0000 /entries/IMP_Compiler_Reuse.html Real-Time Double-Ended Queue /entries/Real_Time_Deque.html Thu, 23 Jun 2022 00:00:00 +0000 /entries/Real_Time_Deque.html Boolos's Curious Inference in Isabelle/HOL /entries/Boolos_Curious_Inference.html Mon, 20 Jun 2022 00:00:00 +0000 /entries/Boolos_Curious_Inference.html Finite Fields /entries/Finite_Fields.html Wed, 08 Jun 2022 00:00:00 +0000 /entries/Finite_Fields.html IsaNet: Formalization of a Verification Framework for Secure Data Plane Protocols /entries/IsaNet.html Wed, 08 Jun 2022 00:00:00 +0000 /entries/IsaNet.html Diophantine Equations and the DPRM Theorem /entries/DPRM_Theorem.html Mon, 06 Jun 2022 00:00:00 +0000 /entries/DPRM_Theorem.html Reducing Rewrite Properties to Properties on Ground Terms /entries/Rewrite_Properties_Reduction.html Thu, 02 Jun 2022 00:00:00 +0000 /entries/Rewrite_Properties_Reduction.html A Restricted Definition of the Magic Wand to Soundly Combine Fractions of a Wand /entries/Combinable_Wands.html Mon, 30 May 2022 00:00:00 +0000 /entries/Combinable_Wands.html The Plünnecke-Ruzsa Inequality /entries/Pluennecke_Ruzsa_Inequality.html Thu, 26 May 2022 00:00:00 +0000 /entries/Pluennecke_Ruzsa_Inequality.html Formalization of a Framework for the Sound Automation of Magic Wands /entries/Package_logic.html Wed, 18 May 2022 00:00:00 +0000 /entries/Package_logic.html Clique is not solvable by monotone circuits of polynomial size /entries/Clique_and_Monotone_Circuits.html Sun, 08 May 2022 00:00:00 +0000 /entries/Clique_and_Monotone_Circuits.html Fisher's Inequality: Linear Algebraic Proof Techniques for Combinatorics /entries/Fishers_Inequality.html Thu, 21 Apr 2022 00:00:00 +0000 /entries/Fishers_Inequality.html Digit Expansions /entries/Digit_Expansions.html Wed, 20 Apr 2022 00:00:00 +0000 /entries/Digit_Expansions.html The Generalized Multiset Ordering is NP-Complete /entries/Multiset_Ordering_NPC.html Wed, 20 Apr 2022 00:00:00 +0000 /entries/Multiset_Ordering_NPC.html The Sophomore's Dream /entries/Sophomores_Dream.html Sun, 10 Apr 2022 00:00:00 +0000 /entries/Sophomores_Dream.html A Combinator Library for Prefix-Free Codes /entries/Prefix_Free_Code_Combinators.html Fri, 08 Apr 2022 00:00:00 +0000 /entries/Prefix_Free_Code_Combinators.html Formalization of Randomized Approximation Algorithms for Frequency Moments /entries/Frequency_Moments.html Fri, 08 Apr 2022 00:00:00 +0000 /entries/Frequency_Moments.html Constructing the Reals as Dedekind Cuts of Rationals /entries/Dedekind_Real.html Thu, 24 Mar 2022 00:00:00 +0000 /entries/Dedekind_Real.html Ackermann's Function Is Not Primitive Recursive /entries/Ackermanns_not_PR.html Wed, 23 Mar 2022 00:00:00 +0000 /entries/Ackermanns_not_PR.html A Naive Prover for First-Order Logic /entries/FOL_Seq_Calc3.html Tue, 22 Mar 2022 00:00:00 +0000 /entries/FOL_Seq_Calc3.html A Proof from THE BOOK: The Partial Fraction Expansion of the Cotangent /entries/Cotangent_PFD_Formula.html Tue, 15 Mar 2022 00:00:00 +0000 /entries/Cotangent_PFD_Formula.html The Independence of the Continuum Hypothesis in Isabelle/ZF /entries/Independence_CH.html Sun, 06 Mar 2022 00:00:00 +0000 /entries/Independence_CH.html Transitive Models of Fragments of ZFC /entries/Transitive_Models.html Thu, 03 Mar 2022 00:00:00 +0000 /entries/Transitive_Models.html Residuated Transition Systems /entries/ResiduatedTransitionSystem.html Mon, 28 Feb 2022 00:00:00 +0000 /entries/ResiduatedTransitionSystem.html Universal Hash Families /entries/Universal_Hash_Families.html Sun, 20 Feb 2022 00:00:00 +0000 /entries/Universal_Hash_Families.html Wetzel's Problem and the Continuum Hypothesis /entries/Wetzels_Problem.html Fri, 18 Feb 2022 00:00:00 +0000 /entries/Wetzels_Problem.html First-Order Query Evaluation /entries/Eval_FO.html Tue, 15 Feb 2022 00:00:00 +0000 /entries/Eval_FO.html Multi-Head Monitoring of Metric Dynamic Logic /entries/VYDRA_MDL.html Sun, 13 Feb 2022 00:00:00 +0000 /entries/VYDRA_MDL.html Enumeration of Equivalence Relations /entries/Equivalence_Relation_Enumeration.html Fri, 04 Feb 2022 00:00:00 +0000 /entries/Equivalence_Relation_Enumeration.html Duality of Linear Programming /entries/LP_Duality.html Thu, 03 Feb 2022 00:00:00 +0000 /entries/LP_Duality.html Quasi-Borel Spaces /entries/Quasi_Borel_Spaces.html Thu, 03 Feb 2022 00:00:00 +0000 /entries/Quasi_Borel_Spaces.html First-Order Theory of Rewriting /entries/FO_Theory_Rewriting.html Wed, 02 Feb 2022 00:00:00 +0000 /entries/FO_Theory_Rewriting.html A Sequent Calculus Prover for First-Order Logic with Functions /entries/FOL_Seq_Calc2.html Mon, 31 Jan 2022 00:00:00 +0000 /entries/FOL_Seq_Calc2.html Young's Inequality for Increasing Functions /entries/Youngs_Inequality.html Mon, 31 Jan 2022 00:00:00 +0000 /entries/Youngs_Inequality.html Interpolation Polynomials (in HOL-Algebra) /entries/Interpolation_Polynomials_HOL_Algebra.html Sat, 29 Jan 2022 00:00:00 +0000 /entries/Interpolation_Polynomials_HOL_Algebra.html Median Method /entries/Median_Method.html Tue, 25 Jan 2022 00:00:00 +0000 /entries/Median_Method.html Actuarial Mathematics /entries/Actuarial_Mathematics.html Sun, 23 Jan 2022 00:00:00 +0000 /entries/Actuarial_Mathematics.html Irrational numbers from THE BOOK /entries/Irrationals_From_THEBOOK.html Sat, 08 Jan 2022 00:00:00 +0000 /entries/Irrationals_From_THEBOOK.html Knight's Tour Revisited Revisited /entries/Knights_Tour.html Tue, 04 Jan 2022 00:00:00 +0000 /entries/Knights_Tour.html Hyperdual Numbers and Forward Differentiation /entries/Hyperdual.html Fri, 31 Dec 2021 00:00:00 +0000 /entries/Hyperdual.html Gale-Shapley Algorithm /entries/Gale_Shapley.html Wed, 29 Dec 2021 00:00:00 +0000 /entries/Gale_Shapley.html Roth's Theorem on Arithmetic Progressions /entries/Roth_Arithmetic_Progressions.html Tue, 28 Dec 2021 00:00:00 +0000 /entries/Roth_Arithmetic_Progressions.html Markov Decision Processes with Rewards /entries/MDP-Rewards.html Thu, 16 Dec 2021 00:00:00 +0000 /entries/MDP-Rewards.html Verified Algorithms for Solving Markov Decision Processes /entries/MDP-Algorithms.html Thu, 16 Dec 2021 00:00:00 +0000 /entries/MDP-Algorithms.html Regular Tree Relations /entries/Regular_Tree_Relations.html Wed, 15 Dec 2021 00:00:00 +0000 /entries/Regular_Tree_Relations.html Simplicial Complexes and Boolean functions /entries/Simplicial_complexes_and_boolean_functions.html Mon, 29 Nov 2021 00:00:00 +0000 /entries/Simplicial_complexes_and_boolean_functions.html van Emde Boas Trees /entries/Van_Emde_Boas_Trees.html Tue, 23 Nov 2021 00:00:00 +0000 /entries/Van_Emde_Boas_Trees.html Foundation of geometry in planes, and some complements: Excluding the parallel axioms /entries/Foundation_of_geometry.html Mon, 22 Nov 2021 00:00:00 +0000 /entries/Foundation_of_geometry.html The Hahn and Jordan Decomposition Theorems /entries/Hahn_Jordan_Decomposition.html Fri, 19 Nov 2021 00:00:00 +0000 /entries/Hahn_Jordan_Decomposition.html Automating Public Announcement Logic and the Wise Men Puzzle in Isabelle/HOL /entries/PAL.html Mon, 08 Nov 2021 00:00:00 +0000 /entries/PAL.html Exploring Simplified Variants of Gödel’s Ontological Argument in Isabelle/HOL /entries/SimplifiedOntologicalArgument.html Mon, 08 Nov 2021 00:00:00 +0000 /entries/SimplifiedOntologicalArgument.html Factorization of Polynomials with Algebraic Coefficients /entries/Factor_Algebraic_Polynomial.html Mon, 08 Nov 2021 00:00:00 +0000 /entries/Factor_Algebraic_Polynomial.html Real Exponents as the Limits of Sequences of Rational Exponents /entries/Real_Power.html Mon, 08 Nov 2021 00:00:00 +0000 /entries/Real_Power.html Szemerédi's Regularity Lemma /entries/Szemeredi_Regularity.html Fri, 05 Nov 2021 00:00:00 +0000 /entries/Szemeredi_Regularity.html Quantum and Classical Registers /entries/Registers.html Thu, 28 Oct 2021 00:00:00 +0000 /entries/Registers.html Belief Revision Theory /entries/Belief_Revision.html Tue, 19 Oct 2021 00:00:00 +0000 /entries/Belief_Revision.html X86 instruction semantics and basic block symbolic execution /entries/X86_Semantics.html Wed, 13 Oct 2021 00:00:00 +0000 /entries/X86_Semantics.html Algebras for Iteration, Infinite Executions and Correctness of Sequential Computations /entries/Correctness_Algebras.html Tue, 12 Oct 2021 00:00:00 +0000 /entries/Correctness_Algebras.html Verified Quadratic Virtual Substitution for Real Arithmetic /entries/Virtual_Substitution.html Sat, 02 Oct 2021 00:00:00 +0000 /entries/Virtual_Substitution.html Soundness and Completeness of an Axiomatic System for First-Order Logic /entries/FOL_Axiomatic.html Fri, 24 Sep 2021 00:00:00 +0000 /entries/FOL_Axiomatic.html Complex Bounded Operators /entries/Complex_Bounded_Operators.html Sat, 18 Sep 2021 00:00:00 +0000 /entries/Complex_Bounded_Operators.html A Formalization of Weighted Path Orders and Recursive Path Orders /entries/Weighted_Path_Order.html Thu, 16 Sep 2021 00:00:00 +0000 /entries/Weighted_Path_Order.html Category Theory for ZFC in HOL I: Foundations: Design Patterns, Set Theory, Digraphs, Semicategories /entries/CZH_Foundations.html Mon, 06 Sep 2021 00:00:00 +0000 /entries/CZH_Foundations.html Category Theory for ZFC in HOL II: Elementary Theory of 1-Categories /entries/CZH_Elementary_Categories.html Mon, 06 Sep 2021 00:00:00 +0000 /entries/CZH_Elementary_Categories.html Category Theory for ZFC in HOL III: Universal Constructions /entries/CZH_Universal_Constructions.html Mon, 06 Sep 2021 00:00:00 +0000 /entries/CZH_Universal_Constructions.html Conditional Simplification /entries/Conditional_Simplification.html Mon, 06 Sep 2021 00:00:00 +0000 /entries/Conditional_Simplification.html Conditional Transfer Rule /entries/Conditional_Transfer_Rule.html Mon, 06 Sep 2021 00:00:00 +0000 /entries/Conditional_Transfer_Rule.html Extension of Types-To-Sets /entries/Types_To_Sets_Extension.html Mon, 06 Sep 2021 00:00:00 +0000 /entries/Types_To_Sets_Extension.html IDE: Introduction, Destruction, Elimination /entries/Intro_Dest_Elim.html Mon, 06 Sep 2021 00:00:00 +0000 /entries/Intro_Dest_Elim.html A data flow analysis algorithm for computing dominators /entries/Dominance_CHK.html Sun, 05 Sep 2021 00:00:00 +0000 /entries/Dominance_CHK.html Solving Cubic and Quartic Equations /entries/Cubic_Quartic_Equations.html Fri, 03 Sep 2021 00:00:00 +0000 /entries/Cubic_Quartic_Equations.html Logging-independent Message Anonymity in the Relational Method /entries/Logging_Independent_Anonymity.html Thu, 26 Aug 2021 00:00:00 +0000 /entries/Logging_Independent_Anonymity.html The Theorem of Three Circles /entries/Three_Circles.html Sat, 21 Aug 2021 00:00:00 +0000 /entries/Three_Circles.html CoCon: A Confidentiality-Verified Conference Management System /entries/CoCon.html Mon, 16 Aug 2021 00:00:00 +0000 /entries/CoCon.html Compositional BD Security /entries/BD_Security_Compositional.html Mon, 16 Aug 2021 00:00:00 +0000 /entries/BD_Security_Compositional.html CoSMed: A confidentiality-verified social media platform /entries/CoSMed.html Mon, 16 Aug 2021 00:00:00 +0000 /entries/CoSMed.html CoSMeDis: A confidentiality-verified distributed social media platform /entries/CoSMeDis.html Mon, 16 Aug 2021 00:00:00 +0000 /entries/CoSMeDis.html Fresh identifiers /entries/Fresh_Identifiers.html Mon, 16 Aug 2021 00:00:00 +0000 /entries/Fresh_Identifiers.html Combinatorial Design Theory /entries/Design_Theory.html Fri, 13 Aug 2021 00:00:00 +0000 /entries/Design_Theory.html Relational Forests /entries/Relational_Forests.html Tue, 03 Aug 2021 00:00:00 +0000 /entries/Relational_Forests.html Schutz' Independent Axioms for Minkowski Spacetime /entries/Schutz_Spacetime.html Tue, 27 Jul 2021 00:00:00 +0000 /entries/Schutz_Spacetime.html Finitely Generated Abelian Groups /entries/Finitely_Generated_Abelian_Groups.html Wed, 07 Jul 2021 00:00:00 +0000 /entries/Finitely_Generated_Abelian_Groups.html SpecCheck - Specification-Based Testing for Isabelle/ML /entries/SpecCheck.html Thu, 01 Jul 2021 00:00:00 +0000 /entries/SpecCheck.html Van der Waerden's Theorem /entries/Van_der_Waerden.html Tue, 22 Jun 2021 00:00:00 +0000 /entries/Van_der_Waerden.html MiniSail - A kernel language for the ISA specification language SAIL /entries/MiniSail.html Fri, 18 Jun 2021 00:00:00 +0000 /entries/MiniSail.html Public Announcement Logic /entries/Public_Announcement_Logic.html Thu, 17 Jun 2021 00:00:00 +0000 /entries/Public_Announcement_Logic.html A Shorter Compiler Correctness Proof for Language IMP /entries/IMP_Compiler.html Fri, 04 Jun 2021 00:00:00 +0000 /entries/IMP_Compiler.html Combinatorics on Words Basics /entries/Combinatorics_Words.html Mon, 24 May 2021 00:00:00 +0000 /entries/Combinatorics_Words.html Graph Lemma /entries/Combinatorics_Words_Graph_Lemma.html Mon, 24 May 2021 00:00:00 +0000 /entries/Combinatorics_Words_Graph_Lemma.html Lyndon words /entries/Combinatorics_Words_Lyndon.html Mon, 24 May 2021 00:00:00 +0000 /entries/Combinatorics_Words_Lyndon.html Regression Test Selection /entries/Regression_Test_Selection.html Fri, 30 Apr 2021 00:00:00 +0000 /entries/Regression_Test_Selection.html Isabelle's Metalogic: Formalization and Proof Checker /entries/Metalogic_ProofChecker.html Tue, 27 Apr 2021 00:00:00 +0000 /entries/Metalogic_ProofChecker.html Lifting the Exponent /entries/Lifting_the_Exponent.html Tue, 27 Apr 2021 00:00:00 +0000 /entries/Lifting_the_Exponent.html The BKR Decision Procedure for Univariate Real Arithmetic /entries/BenOr_Kozen_Reif.html Sat, 24 Apr 2021 00:00:00 +0000 /entries/BenOr_Kozen_Reif.html Gale-Stewart Games /entries/GaleStewart_Games.html Fri, 23 Apr 2021 00:00:00 +0000 /entries/GaleStewart_Games.html Formalization of Timely Dataflow's Progress Tracking Protocol /entries/Progress_Tracking.html Tue, 13 Apr 2021 00:00:00 +0000 /entries/Progress_Tracking.html Information Flow Control via Dependency Tracking /entries/IFC_Tracking.html Thu, 01 Apr 2021 00:00:00 +0000 /entries/IFC_Tracking.html Grothendieck's Schemes in Algebraic Geometry /entries/Grothendieck_Schemes.html Mon, 29 Mar 2021 00:00:00 +0000 /entries/Grothendieck_Schemes.html Hensel's Lemma for the p-adic Integers /entries/Padic_Ints.html Tue, 23 Mar 2021 00:00:00 +0000 /entries/Padic_Ints.html Constructive Cryptography in HOL: the Communication Modeling Aspect /entries/Constructive_Cryptography_CM.html Wed, 17 Mar 2021 00:00:00 +0000 /entries/Constructive_Cryptography_CM.html Two algorithms based on modular arithmetic: lattice basis reduction and Hermite normal form computation /entries/Modular_arithmetic_LLL_and_HNF_algorithms.html Fri, 12 Mar 2021 00:00:00 +0000 /entries/Modular_arithmetic_LLL_and_HNF_algorithms.html Quantum projective measurements and the CHSH inequality /entries/Projective_Measurements.html Wed, 03 Mar 2021 00:00:00 +0000 /entries/Projective_Measurements.html The Hermite–Lindemann–Weierstraß Transcendence Theorem /entries/Hermite_Lindemann.html Wed, 03 Mar 2021 00:00:00 +0000 /entries/Hermite_Lindemann.html Mereology /entries/Mereology.html Mon, 01 Mar 2021 00:00:00 +0000 /entries/Mereology.html The Sunflower Lemma of Erdős and Rado /entries/Sunflowers.html Thu, 25 Feb 2021 00:00:00 +0000 /entries/Sunflowers.html A Verified Imperative Implementation of B-Trees /entries/BTree.html Wed, 24 Feb 2021 00:00:00 +0000 /entries/BTree.html Formal Puiseux Series /entries/Formal_Puiseux_Series.html Wed, 17 Feb 2021 00:00:00 +0000 /entries/Formal_Puiseux_Series.html The Laws of Large Numbers /entries/Laws_of_Large_Numbers.html Wed, 10 Feb 2021 00:00:00 +0000 /entries/Laws_of_Large_Numbers.html Tarski's Parallel Postulate implies the 5th Postulate of Euclid, the Postulate of Playfair and the original Parallel Postulate of Euclid /entries/IsaGeoCoq.html Sun, 31 Jan 2021 00:00:00 +0000 /entries/IsaGeoCoq.html Solution to the xkcd Blue Eyes puzzle /entries/Blue_Eyes.html Sat, 30 Jan 2021 00:00:00 +0000 /entries/Blue_Eyes.html Hood-Melville Queue /entries/Hood_Melville_Queue.html Mon, 18 Jan 2021 00:00:00 +0000 /entries/Hood_Melville_Queue.html JinjaDCI: a Java semantics with dynamic class initialization /entries/JinjaDCI.html Mon, 11 Jan 2021 00:00:00 +0000 /entries/JinjaDCI.html Cofinality and the Delta System Lemma /entries/Delta_System_Lemma.html Sun, 27 Dec 2020 00:00:00 +0000 /entries/Delta_System_Lemma.html Topological semantics for paraconsistent and paracomplete logics /entries/Topological_Semantics.html Thu, 17 Dec 2020 00:00:00 +0000 /entries/Topological_Semantics.html Relational Minimum Spanning Tree Algorithms /entries/Relational_Minimum_Spanning_Trees.html Tue, 08 Dec 2020 00:00:00 +0000 /entries/Relational_Minimum_Spanning_Trees.html Inline Caching and Unboxing Optimization for Interpreters /entries/Interpreter_Optimizations.html Mon, 07 Dec 2020 00:00:00 +0000 /entries/Interpreter_Optimizations.html The Relational Method with Message Anonymity for the Verification of Cryptographic Protocols /entries/Relational_Method.html Sat, 05 Dec 2020 00:00:00 +0000 /entries/Relational_Method.html Isabelle Marries Dirac: a Library for Quantum Computation and Quantum Information /entries/Isabelle_Marries_Dirac.html Sun, 22 Nov 2020 00:00:00 +0000 /entries/Isabelle_Marries_Dirac.html The HOL-CSP Refinement Toolkit /entries/CSP_RefTK.html Thu, 19 Nov 2020 00:00:00 +0000 /entries/CSP_RefTK.html AI Planning Languages Semantics /entries/AI_Planning_Languages_Semantics.html Thu, 29 Oct 2020 00:00:00 +0000 /entries/AI_Planning_Languages_Semantics.html Verified SAT-Based AI Planning /entries/Verified_SAT_Based_AI_Planning.html Thu, 29 Oct 2020 00:00:00 +0000 /entries/Verified_SAT_Based_AI_Planning.html A Sound Type System for Physical Quantities, Units, and Measurements /entries/Physical_Quantities.html Tue, 20 Oct 2020 00:00:00 +0000 /entries/Physical_Quantities.html Finite Map Extras /entries/Finite-Map-Extras.html Mon, 12 Oct 2020 00:00:00 +0000 /entries/Finite-Map-Extras.html A Formal Model of the Document Object Model with Shadow Roots /entries/Shadow_DOM.html Mon, 28 Sep 2020 00:00:00 +0000 /entries/Shadow_DOM.html A Formal Model of the Safely Composable Document Object Model with Shadow Roots /entries/Shadow_SC_DOM.html Mon, 28 Sep 2020 00:00:00 +0000 /entries/Shadow_SC_DOM.html A Formalization of Safely Composable Web Components /entries/SC_DOM_Components.html Mon, 28 Sep 2020 00:00:00 +0000 /entries/SC_DOM_Components.html A Formalization of Web Components /entries/DOM_Components.html Mon, 28 Sep 2020 00:00:00 +0000 /entries/DOM_Components.html The Safely Composable DOM /entries/Core_SC_DOM.html Mon, 28 Sep 2020 00:00:00 +0000 /entries/Core_SC_DOM.html An Abstract Formalization of G&ouml;del's Incompleteness Theorems /entries/Goedel_Incompleteness.html Wed, 16 Sep 2020 00:00:00 +0000 /entries/Goedel_Incompleteness.html From Abstract to Concrete G&ouml;del's Incompleteness Theorems&mdash;Part I /entries/Goedel_HFSet_Semantic.html Wed, 16 Sep 2020 00:00:00 +0000 /entries/Goedel_HFSet_Semantic.html From Abstract to Concrete G&ouml;del's Incompleteness Theorems&mdash;Part II /entries/Goedel_HFSet_Semanticless.html Wed, 16 Sep 2020 00:00:00 +0000 /entries/Goedel_HFSet_Semanticless.html Robinson Arithmetic /entries/Robinson_Arithmetic.html Wed, 16 Sep 2020 00:00:00 +0000 /entries/Robinson_Arithmetic.html Syntax-Independent Logic Infrastructure /entries/Syntax_Independent_Logic.html Wed, 16 Sep 2020 00:00:00 +0000 /entries/Syntax_Independent_Logic.html A Formal Model of Extended Finite State Machines /entries/Extended_Finite_State_Machines.html Mon, 07 Sep 2020 00:00:00 +0000 /entries/Extended_Finite_State_Machines.html Inference of Extended Finite State Machines /entries/Extended_Finite_State_Machine_Inference.html Mon, 07 Sep 2020 00:00:00 +0000 /entries/Extended_Finite_State_Machine_Inference.html Practical Algebraic Calculus Checker /entries/PAC_Checker.html Mon, 31 Aug 2020 00:00:00 +0000 /entries/PAC_Checker.html Some classical results in inductive inference of recursive functions /entries/Inductive_Inference.html Mon, 31 Aug 2020 00:00:00 +0000 /entries/Inductive_Inference.html Relational Disjoint-Set Forests /entries/Relational_Disjoint_Set_Forests.html Wed, 26 Aug 2020 00:00:00 +0000 /entries/Relational_Disjoint_Set_Forests.html Extensions to the Comprehensive Framework for Saturation Theorem Proving /entries/Saturation_Framework_Extensions.html Tue, 25 Aug 2020 00:00:00 +0000 /entries/Saturation_Framework_Extensions.html Putting the `K' into Bird's derivation of Knuth-Morris-Pratt string matching /entries/BirdKMP.html Tue, 25 Aug 2020 00:00:00 +0000 /entries/BirdKMP.html Amicable Numbers /entries/Amicable_Numbers.html Tue, 04 Aug 2020 00:00:00 +0000 /entries/Amicable_Numbers.html Ordinal Partitions /entries/Ordinal_Partitions.html Mon, 03 Aug 2020 00:00:00 +0000 /entries/Ordinal_Partitions.html A Formal Proof of The Chandy--Lamport Distributed Snapshot Algorithm /entries/Chandy_Lamport.html Tue, 21 Jul 2020 00:00:00 +0000 /entries/Chandy_Lamport.html Relational Characterisations of Paths /entries/Relational_Paths.html Mon, 13 Jul 2020 00:00:00 +0000 /entries/Relational_Paths.html A Formally Verified Checker of the Safe Distance Traffic Rules for Autonomous Vehicles /entries/Safe_Distance.html Mon, 01 Jun 2020 00:00:00 +0000 /entries/Safe_Distance.html A verified algorithm for computing the Smith normal form of a matrix /entries/Smith_Normal_Form.html Sat, 23 May 2020 00:00:00 +0000 /entries/Smith_Normal_Form.html The Nash-Williams Partition Theorem /entries/Nash_Williams.html Sat, 16 May 2020 00:00:00 +0000 /entries/Nash_Williams.html A Formalization of Knuth–Bendix Orders /entries/Knuth_Bendix_Order.html Wed, 13 May 2020 00:00:00 +0000 /entries/Knuth_Bendix_Order.html Irrationality Criteria for Series by Erdős and Straus /entries/Irrational_Series_Erdos_Straus.html Tue, 12 May 2020 00:00:00 +0000 /entries/Irrational_Series_Erdos_Straus.html Recursion Theorem in ZF /entries/Recursion-Addition.html Mon, 11 May 2020 00:00:00 +0000 /entries/Recursion-Addition.html An Efficient Normalisation Procedure for Linear Temporal Logic: Isabelle/HOL Formalisation /entries/LTL_Normal_Form.html Fri, 08 May 2020 00:00:00 +0000 /entries/LTL_Normal_Form.html Formalization of Forcing in Isabelle/ZF /entries/Forcing.html Wed, 06 May 2020 00:00:00 +0000 /entries/Forcing.html Banach-Steinhaus Theorem /entries/Banach_Steinhaus.html Sat, 02 May 2020 00:00:00 +0000 /entries/Banach_Steinhaus.html Attack Trees in Isabelle for GDPR compliance of IoT healthcare systems /entries/Attack_Trees.html Mon, 27 Apr 2020 00:00:00 +0000 /entries/Attack_Trees.html Gaussian Integers /entries/Gaussian_Integers.html Fri, 24 Apr 2020 00:00:00 +0000 /entries/Gaussian_Integers.html Power Sum Polynomials /entries/Power_Sum_Polynomials.html Fri, 24 Apr 2020 00:00:00 +0000 /entries/Power_Sum_Polynomials.html The Lambert W Function on the Reals /entries/Lambert_W.html Fri, 24 Apr 2020 00:00:00 +0000 /entries/Lambert_W.html Matrices for ODEs /entries/Matrices_for_ODEs.html Sun, 19 Apr 2020 00:00:00 +0000 /entries/Matrices_for_ODEs.html Authenticated Data Structures As Functors /entries/ADS_Functor.html Thu, 16 Apr 2020 00:00:00 +0000 /entries/ADS_Functor.html Formalization of an Algorithm for Greedily Computing Associative Aggregations on Sliding Windows /entries/Sliding_Window_Algorithm.html Fri, 10 Apr 2020 00:00:00 +0000 /entries/Sliding_Window_Algorithm.html A Comprehensive Framework for Saturation Theorem Proving /entries/Saturation_Framework.html Thu, 09 Apr 2020 00:00:00 +0000 /entries/Saturation_Framework.html Formalization of an Optimized Monitoring Algorithm for Metric First-Order Dynamic Logic with Aggregations /entries/MFODL_Monitor_Optimized.html Thu, 09 Apr 2020 00:00:00 +0000 /entries/MFODL_Monitor_Optimized.html Automated Stateful Protocol Verification /entries/Automated_Stateful_Protocol_Verification.html Wed, 08 Apr 2020 00:00:00 +0000 /entries/Automated_Stateful_Protocol_Verification.html Stateful Protocol Composition and Typing /entries/Stateful_Protocol_Composition_and_Typing.html Wed, 08 Apr 2020 00:00:00 +0000 /entries/Stateful_Protocol_Composition_and_Typing.html Lucas's Theorem /entries/Lucas_Theorem.html Tue, 07 Apr 2020 00:00:00 +0000 /entries/Lucas_Theorem.html Strong Eventual Consistency of the Collaborative Editing Framework WOOT /entries/WOOT_Strong_Eventual_Consistency.html Wed, 25 Mar 2020 00:00:00 +0000 /entries/WOOT_Strong_Eventual_Consistency.html Furstenberg's topology and his proof of the infinitude of primes /entries/Furstenberg_Topology.html Sun, 22 Mar 2020 00:00:00 +0000 /entries/Furstenberg_Topology.html An Under-Approximate Relational Logic /entries/Relational-Incorrectness-Logic.html Thu, 12 Mar 2020 00:00:00 +0000 /entries/Relational-Incorrectness-Logic.html Hello World /entries/Hello_World.html Sat, 07 Mar 2020 00:00:00 +0000 /entries/Hello_World.html Implementing the Goodstein Function in &lambda;-Calculus /entries/Goodstein_Lambda.html Fri, 21 Feb 2020 00:00:00 +0000 /entries/Goodstein_Lambda.html A Generic Framework for Verified Compilers /entries/VeriComp.html Mon, 10 Feb 2020 00:00:00 +0000 /entries/VeriComp.html Arithmetic progressions and relative primes /entries/Arith_Prog_Rel_Primes.html Sat, 01 Feb 2020 00:00:00 +0000 /entries/Arith_Prog_Rel_Primes.html A Hierarchy of Algebras for Boolean Subsets /entries/Subset_Boolean_Algebras.html Fri, 31 Jan 2020 00:00:00 +0000 /entries/Subset_Boolean_Algebras.html Mersenne primes and the Lucas–Lehmer test /entries/Mersenne_Primes.html Fri, 17 Jan 2020 00:00:00 +0000 /entries/Mersenne_Primes.html Verified Approximation Algorithms /entries/Approximation_Algorithms.html Thu, 16 Jan 2020 00:00:00 +0000 /entries/Approximation_Algorithms.html Closest Pair of Points Algorithms /entries/Closest_Pair_Points.html Mon, 13 Jan 2020 00:00:00 +0000 /entries/Closest_Pair_Points.html Skip Lists /entries/Skip_Lists.html Thu, 09 Jan 2020 00:00:00 +0000 /entries/Skip_Lists.html Bicategories /entries/Bicategory.html Mon, 06 Jan 2020 00:00:00 +0000 /entries/Bicategory.html The Irrationality of ζ(3) /entries/Zeta_3_Irrational.html Fri, 27 Dec 2019 00:00:00 +0000 /entries/Zeta_3_Irrational.html Formalizing a Seligman-Style Tableau System for Hybrid Logic /entries/Hybrid_Logic.html Fri, 20 Dec 2019 00:00:00 +0000 /entries/Hybrid_Logic.html The Poincaré-Bendixson Theorem /entries/Poincare_Bendixson.html Wed, 18 Dec 2019 00:00:00 +0000 /entries/Poincare_Bendixson.html Complex Geometry /entries/Complex_Geometry.html Mon, 16 Dec 2019 00:00:00 +0000 /entries/Complex_Geometry.html Poincaré Disc Model /entries/Poincare_Disc.html Mon, 16 Dec 2019 00:00:00 +0000 /entries/Poincare_Disc.html Gauss Sums and the Pólya–Vinogradov Inequality /entries/Gauss_Sums.html Tue, 10 Dec 2019 00:00:00 +0000 /entries/Gauss_Sums.html An Efficient Generalization of Counting Sort for Large, possibly Infinite Key Ranges /entries/Generalized_Counting_Sort.html Wed, 04 Dec 2019 00:00:00 +0000 /entries/Generalized_Counting_Sort.html Interval Arithmetic on 32-bit Words /entries/Interval_Arithmetic_Word32.html Wed, 27 Nov 2019 00:00:00 +0000 /entries/Interval_Arithmetic_Word32.html Zermelo Fraenkel Set Theory in Higher-Order Logic /entries/ZFC_in_HOL.html Thu, 24 Oct 2019 00:00:00 +0000 /entries/ZFC_in_HOL.html Isabelle/C /entries/Isabelle_C.html Tue, 22 Oct 2019 00:00:00 +0000 /entries/Isabelle_C.html VerifyThis 2019 -- Polished Isabelle Solutions /entries/VerifyThis2019.html Wed, 16 Oct 2019 00:00:00 +0000 /entries/VerifyThis2019.html Aristotle's Assertoric Syllogistic /entries/Aristotles_Assertoric_Syllogistic.html Tue, 08 Oct 2019 00:00:00 +0000 /entries/Aristotles_Assertoric_Syllogistic.html Sigma Protocols and Commitment Schemes /entries/Sigma_Commit_Crypto.html Mon, 07 Oct 2019 00:00:00 +0000 /entries/Sigma_Commit_Crypto.html Clean - An Abstract Imperative Programming Language and its Theory /entries/Clean.html Fri, 04 Oct 2019 00:00:00 +0000 /entries/Clean.html Formalization of Multiway-Join Algorithms /entries/Generic_Join.html Mon, 16 Sep 2019 00:00:00 +0000 /entries/Generic_Join.html Verification Components for Hybrid Systems /entries/Hybrid_Systems_VCs.html Tue, 10 Sep 2019 00:00:00 +0000 /entries/Hybrid_Systems_VCs.html Fourier Series /entries/Fourier.html Fri, 06 Sep 2019 00:00:00 +0000 /entries/Fourier.html A Case Study in Basic Algebra /entries/Jacobson_Basic_Algebra.html Fri, 30 Aug 2019 00:00:00 +0000 /entries/Jacobson_Basic_Algebra.html Formalisation of an Adaptive State Counting Algorithm /entries/Adaptive_State_Counting.html Fri, 16 Aug 2019 00:00:00 +0000 /entries/Adaptive_State_Counting.html Laplace Transform /entries/Laplace_Transform.html Wed, 14 Aug 2019 00:00:00 +0000 /entries/Laplace_Transform.html Communicating Concurrent Kleene Algebra for Distributed Systems Specification /entries/C2KA_DistributedSystems.html Tue, 06 Aug 2019 00:00:00 +0000 /entries/C2KA_DistributedSystems.html Linear Programming /entries/Linear_Programming.html Tue, 06 Aug 2019 00:00:00 +0000 /entries/Linear_Programming.html Selected Problems from the International Mathematical Olympiad 2019 /entries/IMO2019.html Mon, 05 Aug 2019 00:00:00 +0000 /entries/IMO2019.html Stellar Quorum Systems /entries/Stellar_Quorums.html Thu, 01 Aug 2019 00:00:00 +0000 /entries/Stellar_Quorums.html A Formal Development of a Polychronous Polytimed Coordination Language /entries/TESL_Language.html Tue, 30 Jul 2019 00:00:00 +0000 /entries/TESL_Language.html Order Extension and Szpilrajn's Extension Theorem /entries/Szpilrajn.html Sat, 27 Jul 2019 00:00:00 +0000 /entries/Szpilrajn.html A Sequent Calculus for First-Order Logic /entries/FOL_Seq_Calc1.html Thu, 18 Jul 2019 00:00:00 +0000 /entries/FOL_Seq_Calc1.html A Verified Code Generator from Isabelle/HOL to CakeML /entries/CakeML_Codegen.html Mon, 08 Jul 2019 00:00:00 +0000 /entries/CakeML_Codegen.html Formalization of a Monitoring Algorithm for Metric First-Order Temporal Logic /entries/MFOTL_Monitor.html Thu, 04 Jul 2019 00:00:00 +0000 /entries/MFOTL_Monitor.html Complete Non-Orders and Fixed Points /entries/Complete_Non_Orders.html Thu, 27 Jun 2019 00:00:00 +0000 /entries/Complete_Non_Orders.html Priority Search Trees /entries/Priority_Search_Trees.html Tue, 25 Jun 2019 00:00:00 +0000 /entries/Priority_Search_Trees.html Purely Functional, Simple, and Efficient Implementation of Prim and Dijkstra /entries/Prim_Dijkstra_Simple.html Tue, 25 Jun 2019 00:00:00 +0000 /entries/Prim_Dijkstra_Simple.html Linear Inequalities /entries/Linear_Inequalities.html Fri, 21 Jun 2019 00:00:00 +0000 /entries/Linear_Inequalities.html Hilbert's Nullstellensatz /entries/Nullstellensatz.html Sun, 16 Jun 2019 00:00:00 +0000 /entries/Nullstellensatz.html Gröbner Bases, Macaulay Matrices and Dubé's Degree Bounds /entries/Groebner_Macaulay.html Sat, 15 Jun 2019 00:00:00 +0000 /entries/Groebner_Macaulay.html Binary Heaps for IMP2 /entries/IMP2_Binary_Heap.html Thu, 13 Jun 2019 00:00:00 +0000 /entries/IMP2_Binary_Heap.html Differential Game Logic /entries/Differential_Game_Logic.html Mon, 03 Jun 2019 00:00:00 +0000 /entries/Differential_Game_Logic.html Multidimensional Binary Search Trees /entries/KD_Tree.html Thu, 30 May 2019 00:00:00 +0000 /entries/KD_Tree.html Formalization of Generic Authenticated Data Structures /entries/LambdaAuth.html Tue, 14 May 2019 00:00:00 +0000 /entries/LambdaAuth.html Multi-Party Computation /entries/Multi_Party_Computation.html Thu, 09 May 2019 00:00:00 +0000 /entries/Multi_Party_Computation.html HOL-CSP Version 2.0 /entries/HOL-CSP.html Fri, 26 Apr 2019 00:00:00 +0000 /entries/HOL-CSP.html A Compositional and Unified Translation of LTL into ω-Automata /entries/LTL_Master_Theorem.html Tue, 16 Apr 2019 00:00:00 +0000 /entries/LTL_Master_Theorem.html A General Theory of Syntax with Bindings /entries/Binding_Syntax_Theory.html Sat, 06 Apr 2019 00:00:00 +0000 /entries/Binding_Syntax_Theory.html The Transcendence of Certain Infinite Series /entries/Transcendence_Series_Hancl_Rucki.html Wed, 27 Mar 2019 00:00:00 +0000 /entries/Transcendence_Series_Hancl_Rucki.html Quantum Hoare Logic /entries/QHLProver.html Sun, 24 Mar 2019 00:00:00 +0000 /entries/QHLProver.html Safe OCL /entries/Safe_OCL.html Sat, 09 Mar 2019 00:00:00 +0000 /entries/Safe_OCL.html Elementary Facts About the Distribution of Primes /entries/Prime_Distribution_Elementary.html Thu, 21 Feb 2019 00:00:00 +0000 /entries/Prime_Distribution_Elementary.html Kruskal's Algorithm for Minimum Spanning Forest /entries/Kruskal.html Thu, 14 Feb 2019 00:00:00 +0000 /entries/Kruskal.html Probabilistic Primality Testing /entries/Probabilistic_Prime_Tests.html Mon, 11 Feb 2019 00:00:00 +0000 /entries/Probabilistic_Prime_Tests.html Universal Turing Machine /entries/Universal_Turing_Machine.html Fri, 08 Feb 2019 00:00:00 +0000 /entries/Universal_Turing_Machine.html Isabelle/UTP: Mechanised Theory Engineering for Unifying Theories of Programming /entries/UTP.html Fri, 01 Feb 2019 00:00:00 +0000 /entries/UTP.html The Inversions of a List /entries/List_Inversions.html Fri, 01 Feb 2019 00:00:00 +0000 /entries/List_Inversions.html Farkas' Lemma and Motzkin's Transposition Theorem /entries/Farkas.html Thu, 17 Jan 2019 00:00:00 +0000 /entries/Farkas.html An Algebra for Higher-Order Terms /entries/Higher_Order_Terms.html Tue, 15 Jan 2019 00:00:00 +0000 /entries/Higher_Order_Terms.html IMP2 – Simple Program Verification in Isabelle/HOL /entries/IMP2.html Tue, 15 Jan 2019 00:00:00 +0000 /entries/IMP2.html A Reduction Theorem for Store Buffers /entries/Store_Buffer_Reduction.html Mon, 07 Jan 2019 00:00:00 +0000 /entries/Store_Buffer_Reduction.html A Formal Model of the Document Object Model /entries/Core_DOM.html Wed, 26 Dec 2018 00:00:00 +0000 /entries/Core_DOM.html Formalization of Concurrent Revisions /entries/Concurrent_Revisions.html Tue, 25 Dec 2018 00:00:00 +0000 /entries/Concurrent_Revisions.html Verifying Imperative Programs using Auto2 /entries/Auto2_Imperative_HOL.html Fri, 21 Dec 2018 00:00:00 +0000 /entries/Auto2_Imperative_HOL.html Constructive Cryptography in HOL /entries/Constructive_Cryptography.html Mon, 17 Dec 2018 00:00:00 +0000 /entries/Constructive_Cryptography.html Properties of Orderings and Lattices /entries/Order_Lattice_Props.html Tue, 11 Dec 2018 00:00:00 +0000 /entries/Order_Lattice_Props.html Quantales /entries/Quantales.html Tue, 11 Dec 2018 00:00:00 +0000 /entries/Quantales.html Transformer Semantics /entries/Transformer_Semantics.html Tue, 11 Dec 2018 00:00:00 +0000 /entries/Transformer_Semantics.html A Verified Functional Implementation of Bachmair and Ganzinger's Ordered Resolution Prover /entries/Functional_Ordered_Resolution_Prover.html Fri, 23 Nov 2018 00:00:00 +0000 /entries/Functional_Ordered_Resolution_Prover.html Graph Saturation /entries/Graph_Saturation.html Fri, 23 Nov 2018 00:00:00 +0000 /entries/Graph_Saturation.html Auto2 Prover /entries/Auto2_HOL.html Tue, 20 Nov 2018 00:00:00 +0000 /entries/Auto2_HOL.html Matroids /entries/Matroids.html Fri, 16 Nov 2018 00:00:00 +0000 /entries/Matroids.html Deriving generic class instances for datatypes /entries/Generic_Deriving.html Tue, 06 Nov 2018 00:00:00 +0000 /entries/Generic_Deriving.html Formalisation and Evaluation of Alan Gewirth's Proof for the Principle of Generic Consistency in Isabelle/HOL /entries/GewirthPGCProof.html Tue, 30 Oct 2018 00:00:00 +0000 /entries/GewirthPGCProof.html Epistemic Logic: Completeness of Modal Logics /entries/Epistemic_Logic.html Mon, 29 Oct 2018 00:00:00 +0000 /entries/Epistemic_Logic.html Smooth Manifolds /entries/Smooth_Manifolds.html Mon, 22 Oct 2018 00:00:00 +0000 /entries/Smooth_Manifolds.html Formalization of the Embedding Path Order for Lambda-Free Higher-Order Terms /entries/Lambda_Free_EPO.html Fri, 19 Oct 2018 00:00:00 +0000 /entries/Lambda_Free_EPO.html Randomised Binary Search Trees /entries/Randomised_BSTs.html Fri, 19 Oct 2018 00:00:00 +0000 /entries/Randomised_BSTs.html Upper Bounding Diameters of State Spaces of Factored Transition Systems /entries/Factored_Transition_System_Bounding.html Fri, 12 Oct 2018 00:00:00 +0000 /entries/Factored_Transition_System_Bounding.html The Transcendence of π /entries/Pi_Transcendental.html Fri, 28 Sep 2018 00:00:00 +0000 /entries/Pi_Transcendental.html Symmetric Polynomials /entries/Symmetric_Polynomials.html Tue, 25 Sep 2018 00:00:00 +0000 /entries/Symmetric_Polynomials.html Signature-Based Gröbner Basis Algorithms /entries/Signature_Groebner.html Thu, 20 Sep 2018 00:00:00 +0000 /entries/Signature_Groebner.html The Prime Number Theorem /entries/Prime_Number_Theorem.html Wed, 19 Sep 2018 00:00:00 +0000 /entries/Prime_Number_Theorem.html Aggregation Algebras /entries/Aggregation_Algebras.html Sat, 15 Sep 2018 00:00:00 +0000 /entries/Aggregation_Algebras.html Octonions /entries/Octonions.html Fri, 14 Sep 2018 00:00:00 +0000 /entries/Octonions.html Quaternions /entries/Quaternions.html Wed, 05 Sep 2018 00:00:00 +0000 /entries/Quaternions.html The Budan-Fourier Theorem and Counting Real Roots with Multiplicity /entries/Budan_Fourier.html Sun, 02 Sep 2018 00:00:00 +0000 /entries/Budan_Fourier.html An Incremental Simplex Algorithm with Unsatisfiable Core Generation /entries/Simplex.html Fri, 24 Aug 2018 00:00:00 +0000 /entries/Simplex.html Minsky Machines /entries/Minsky_Machines.html Tue, 14 Aug 2018 00:00:00 +0000 /entries/Minsky_Machines.html Pricing in discrete financial models /entries/DiscretePricing.html Mon, 16 Jul 2018 00:00:00 +0000 /entries/DiscretePricing.html Von-Neumann-Morgenstern Utility Theorem /entries/Neumann_Morgenstern_Utility.html Wed, 04 Jul 2018 00:00:00 +0000 /entries/Neumann_Morgenstern_Utility.html Pell's Equation /entries/Pell.html Sat, 23 Jun 2018 00:00:00 +0000 /entries/Pell.html Projective Geometry /entries/Projective_Geometry.html Thu, 14 Jun 2018 00:00:00 +0000 /entries/Projective_Geometry.html The Localization of a Commutative Ring /entries/Localization_Ring.html Thu, 14 Jun 2018 00:00:00 +0000 /entries/Localization_Ring.html Partial Order Reduction /entries/Partial_Order_Reduction.html Tue, 05 Jun 2018 00:00:00 +0000 /entries/Partial_Order_Reduction.html Optimal Binary Search Trees /entries/Optimal_BST.html Sun, 27 May 2018 00:00:00 +0000 /entries/Optimal_BST.html Hidden Markov Models /entries/Hidden_Markov_Models.html Fri, 25 May 2018 00:00:00 +0000 /entries/Hidden_Markov_Models.html Probabilistic Timed Automata /entries/Probabilistic_Timed_Automata.html Thu, 24 May 2018 00:00:00 +0000 /entries/Probabilistic_Timed_Automata.html Axiom Systems for Category Theory in Free Logic /entries/AxiomaticCategoryTheory.html Wed, 23 May 2018 00:00:00 +0000 /entries/AxiomaticCategoryTheory.html Irrational Rapidly Convergent Series /entries/Irrationality_J_Hancl.html Wed, 23 May 2018 00:00:00 +0000 /entries/Irrationality_J_Hancl.html Monadification, Memoization and Dynamic Programming /entries/Monad_Memo_DP.html Tue, 22 May 2018 00:00:00 +0000 /entries/Monad_Memo_DP.html OpSets: Sequential Specifications for Replicated Datatypes /entries/OpSets.html Thu, 10 May 2018 00:00:00 +0000 /entries/OpSets.html An Isabelle/HOL Formalization of the Modular Assembly Kit for Security Properties /entries/Modular_Assembly_Kit_Security.html Mon, 07 May 2018 00:00:00 +0000 /entries/Modular_Assembly_Kit_Security.html WebAssembly /entries/WebAssembly.html Sun, 29 Apr 2018 00:00:00 +0000 /entries/WebAssembly.html VerifyThis 2018 - Polished Isabelle Solutions /entries/VerifyThis2018.html Fri, 27 Apr 2018 00:00:00 +0000 /entries/VerifyThis2018.html Bounded Natural Functors with Covariance and Contravariance /entries/BNF_CC.html Tue, 24 Apr 2018 00:00:00 +0000 /entries/BNF_CC.html The Incompatibility of Fishburn-Strategyproofness and Pareto-Efficiency /entries/Fishburn_Impossibility.html Thu, 22 Mar 2018 00:00:00 +0000 /entries/Fishburn_Impossibility.html Weight-Balanced Trees /entries/Weight_Balanced_Trees.html Tue, 13 Mar 2018 00:00:00 +0000 /entries/Weight_Balanced_Trees.html CakeML /entries/CakeML.html Mon, 12 Mar 2018 00:00:00 +0000 /entries/CakeML.html A Theory of Architectural Design Patterns /entries/Architectural_Design_Patterns.html Thu, 01 Mar 2018 00:00:00 +0000 /entries/Architectural_Design_Patterns.html Hoare Logics for Time Bounds /entries/Hoare_Time.html Mon, 26 Feb 2018 00:00:00 +0000 /entries/Hoare_Time.html A verified factorization algorithm for integer polynomials with polynomial complexity /entries/LLL_Factorization.html Tue, 06 Feb 2018 00:00:00 +0000 /entries/LLL_Factorization.html First-Order Terms /entries/First_Order_Terms.html Tue, 06 Feb 2018 00:00:00 +0000 /entries/First_Order_Terms.html The Error Function /entries/Error_Function.html Tue, 06 Feb 2018 00:00:00 +0000 /entries/Error_Function.html Treaps /entries/Treaps.html Tue, 06 Feb 2018 00:00:00 +0000 /entries/Treaps.html A verified LLL algorithm /entries/LLL_Basis_Reduction.html Fri, 02 Feb 2018 00:00:00 +0000 /entries/LLL_Basis_Reduction.html Formalization of Bachmair and Ganzinger's Ordered Resolution Prover /entries/Ordered_Resolution_Prover.html Thu, 18 Jan 2018 00:00:00 +0000 /entries/Ordered_Resolution_Prover.html Gromov Hyperbolicity /entries/Gromov_Hyperbolicity.html Tue, 16 Jan 2018 00:00:00 +0000 /entries/Gromov_Hyperbolicity.html An Isabelle/HOL formalisation of Green's Theorem /entries/Green.html Thu, 11 Jan 2018 00:00:00 +0000 /entries/Green.html Taylor Models /entries/Taylor_Models.html Mon, 08 Jan 2018 00:00:00 +0000 /entries/Taylor_Models.html The Falling Factorial of a Sum /entries/Falling_Factorial_Sum.html Fri, 22 Dec 2017 00:00:00 +0000 /entries/Falling_Factorial_Sum.html Dirichlet L-Functions and Dirichlet's Theorem /entries/Dirichlet_L.html Thu, 21 Dec 2017 00:00:00 +0000 /entries/Dirichlet_L.html The Mason–Stothers Theorem /entries/Mason_Stothers.html Thu, 21 Dec 2017 00:00:00 +0000 /entries/Mason_Stothers.html The Median-of-Medians Selection Algorithm /entries/Median_Of_Medians_Selection.html Thu, 21 Dec 2017 00:00:00 +0000 /entries/Median_Of_Medians_Selection.html Operations on Bounded Natural Functors /entries/BNF_Operations.html Tue, 19 Dec 2017 00:00:00 +0000 /entries/BNF_Operations.html The string search algorithm by Knuth, Morris and Pratt /entries/Knuth_Morris_Pratt.html Mon, 18 Dec 2017 00:00:00 +0000 /entries/Knuth_Morris_Pratt.html Stochastic Matrices and the Perron-Frobenius Theorem /entries/Stochastic_Matrices.html Wed, 22 Nov 2017 00:00:00 +0000 /entries/Stochastic_Matrices.html The IMAP CmRDT /entries/IMAP-CRDT.html Thu, 09 Nov 2017 00:00:00 +0000 /entries/IMAP-CRDT.html Hybrid Multi-Lane Spatial Logic /entries/Hybrid_Multi_Lane_Spatial_Logic.html Mon, 06 Nov 2017 00:00:00 +0000 /entries/Hybrid_Multi_Lane_Spatial_Logic.html The Kuratowski Closure-Complement Theorem /entries/Kuratowski_Closure_Complement.html Thu, 26 Oct 2017 00:00:00 +0000 /entries/Kuratowski_Closure_Complement.html Büchi Complementation /entries/Buchi_Complementation.html Thu, 19 Oct 2017 00:00:00 +0000 /entries/Buchi_Complementation.html Transition Systems and Automata /entries/Transition_Systems_and_Automata.html Thu, 19 Oct 2017 00:00:00 +0000 /entries/Transition_Systems_and_Automata.html Count the Number of Complex Roots /entries/Count_Complex_Roots.html Tue, 17 Oct 2017 00:00:00 +0000 /entries/Count_Complex_Roots.html Evaluate Winding Numbers through Cauchy Indices /entries/Winding_Number_Eval.html Tue, 17 Oct 2017 00:00:00 +0000 /entries/Winding_Number_Eval.html Homogeneous Linear Diophantine Equations /entries/Diophantine_Eqns_Lin_Hom.html Sat, 14 Oct 2017 00:00:00 +0000 /entries/Diophantine_Eqns_Lin_Hom.html Dirichlet Series /entries/Dirichlet_Series.html Thu, 12 Oct 2017 00:00:00 +0000 /entries/Dirichlet_Series.html Linear Recurrences /entries/Linear_Recurrences.html Thu, 12 Oct 2017 00:00:00 +0000 /entries/Linear_Recurrences.html The Hurwitz and Riemann ζ Functions /entries/Zeta_Function.html Thu, 12 Oct 2017 00:00:00 +0000 /entries/Zeta_Function.html Computer-assisted Reconstruction and Assessment of E. J. Lowe's Modal Ontological Argument /entries/Lowe_Ontological_Argument.html Thu, 21 Sep 2017 00:00:00 +0000 /entries/Lowe_Ontological_Argument.html Representation and Partial Automation of the Principia Logico-Metaphysica in Isabelle/HOL /entries/PLM.html Sun, 17 Sep 2017 00:00:00 +0000 /entries/PLM.html Anselm's God in Isabelle/HOL /entries/AnselmGod.html Wed, 06 Sep 2017 00:00:00 +0000 /entries/AnselmGod.html Microeconomics and the First Welfare Theorem /entries/First_Welfare_Theorem.html Fri, 01 Sep 2017 00:00:00 +0000 /entries/First_Welfare_Theorem.html Orbit-Stabiliser Theorem with Application to Rotational Symmetries /entries/Orbit_Stabiliser.html Sun, 20 Aug 2017 00:00:00 +0000 /entries/Orbit_Stabiliser.html Root-Balanced Tree /entries/Root_Balanced_Tree.html Sun, 20 Aug 2017 00:00:00 +0000 /entries/Root_Balanced_Tree.html The LambdaMu-calculus /entries/LambdaMu.html Wed, 16 Aug 2017 00:00:00 +0000 /entries/LambdaMu.html Stewart's Theorem and Apollonius' Theorem /entries/Stewart_Apollonius.html Mon, 31 Jul 2017 00:00:00 +0000 /entries/Stewart_Apollonius.html Dynamic Architectures /entries/DynamicArchitectures.html Fri, 28 Jul 2017 00:00:00 +0000 /entries/DynamicArchitectures.html Declarative Semantics for Functional Languages /entries/Decl_Sem_Fun_PL.html Fri, 21 Jul 2017 00:00:00 +0000 /entries/Decl_Sem_Fun_PL.html HOLCF-Prelude /entries/HOLCF-Prelude.html Sat, 15 Jul 2017 00:00:00 +0000 /entries/HOLCF-Prelude.html Minkowski's Theorem /entries/Minkowskis_Theorem.html Thu, 13 Jul 2017 00:00:00 +0000 /entries/Minkowskis_Theorem.html Verified Metatheory and Type Inference for a Name-Carrying Simply-Typed Lambda Calculus /entries/Name_Carrying_Type_Inference.html Sun, 09 Jul 2017 00:00:00 +0000 /entries/Name_Carrying_Type_Inference.html A framework for establishing Strong Eventual Consistency for Conflict-free Replicated Datatypes /entries/CRDT.html Fri, 07 Jul 2017 00:00:00 +0000 /entries/CRDT.html Stone-Kleene Relation Algebras /entries/Stone_Kleene_Relation_Algebras.html Thu, 06 Jul 2017 00:00:00 +0000 /entries/Stone_Kleene_Relation_Algebras.html Propositional Proof Systems /entries/Propositional_Proof_Systems.html Wed, 21 Jun 2017 00:00:00 +0000 /entries/Propositional_Proof_Systems.html Partial Semigroups and Convolution Algebras /entries/PSemigroupsConvolution.html Tue, 13 Jun 2017 00:00:00 +0000 /entries/PSemigroupsConvolution.html Buffon's Needle Problem /entries/Buffons_Needle.html Tue, 06 Jun 2017 00:00:00 +0000 /entries/Buffons_Needle.html Flow Networks and the Min-Cut-Max-Flow Theorem /entries/Flow_Networks.html Thu, 01 Jun 2017 00:00:00 +0000 /entries/Flow_Networks.html Formalizing Push-Relabel Algorithms /entries/Prpu_Maxflow.html Thu, 01 Jun 2017 00:00:00 +0000 /entries/Prpu_Maxflow.html Optics /entries/Optics.html Thu, 25 May 2017 00:00:00 +0000 /entries/Optics.html Developing Security Protocols by Refinement /entries/Security_Protocol_Refinement.html Wed, 24 May 2017 00:00:00 +0000 /entries/Security_Protocol_Refinement.html Dictionary Construction /entries/Dict_Construction.html Wed, 24 May 2017 00:00:00 +0000 /entries/Dict_Construction.html The Floyd-Warshall Algorithm for Shortest Paths /entries/Floyd_Warshall.html Mon, 08 May 2017 00:00:00 +0000 /entries/Floyd_Warshall.html CryptHOL /entries/CryptHOL.html Fri, 05 May 2017 00:00:00 +0000 /entries/CryptHOL.html Effect polymorphism in higher-order logic /entries/Monomorphic_Monad.html Fri, 05 May 2017 00:00:00 +0000 /entries/Monomorphic_Monad.html Game-based cryptography in HOL /entries/Game_Based_Crypto.html Fri, 05 May 2017 00:00:00 +0000 /entries/Game_Based_Crypto.html Monad normalisation /entries/Monad_Normalisation.html Fri, 05 May 2017 00:00:00 +0000 /entries/Monad_Normalisation.html Probabilistic while loop /entries/Probabilistic_While.html Fri, 05 May 2017 00:00:00 +0000 /entries/Probabilistic_While.html Monoidal Categories /entries/MonoidalCategory.html Thu, 04 May 2017 00:00:00 +0000 /entries/MonoidalCategory.html Types, Tableaus and Gödel’s God in Isabelle/HOL /entries/Types_Tableaus_and_Goedels_God.html Mon, 01 May 2017 00:00:00 +0000 /entries/Types_Tableaus_and_Goedels_God.html Local Lexing /entries/LocalLexing.html Fri, 28 Apr 2017 00:00:00 +0000 /entries/LocalLexing.html Constructor Functions /entries/Constructor_Funs.html Wed, 19 Apr 2017 00:00:00 +0000 /entries/Constructor_Funs.html Lazifying case constants /entries/Lazy_Case.html Tue, 18 Apr 2017 00:00:00 +0000 /entries/Lazy_Case.html Subresultants /entries/Subresultants.html Thu, 06 Apr 2017 00:00:00 +0000 /entries/Subresultants.html Expected Shape of Random Binary Search Trees /entries/Random_BSTs.html Tue, 04 Apr 2017 00:00:00 +0000 /entries/Random_BSTs.html Lower bound on comparison-based sorting algorithms /entries/Comparison_Sort_Lower_Bound.html Wed, 15 Mar 2017 00:00:00 +0000 /entries/Comparison_Sort_Lower_Bound.html The number of comparisons in QuickSort /entries/Quick_Sort_Cost.html Wed, 15 Mar 2017 00:00:00 +0000 /entries/Quick_Sort_Cost.html The Euler–MacLaurin Formula /entries/Euler_MacLaurin.html Fri, 10 Mar 2017 00:00:00 +0000 /entries/Euler_MacLaurin.html The Group Law for Elliptic Curves /entries/Elliptic_Curves_Group_Law.html Tue, 28 Feb 2017 00:00:00 +0000 /entries/Elliptic_Curves_Group_Law.html Menger's Theorem /entries/Menger.html Sun, 26 Feb 2017 00:00:00 +0000 /entries/Menger.html Differential Dynamic Logic /entries/Differential_Dynamic_Logic.html Mon, 13 Feb 2017 00:00:00 +0000 /entries/Differential_Dynamic_Logic.html Abstract Soundness /entries/Abstract_Soundness.html Fri, 10 Feb 2017 00:00:00 +0000 /entries/Abstract_Soundness.html Stone Relation Algebras /entries/Stone_Relation_Algebras.html Tue, 07 Feb 2017 00:00:00 +0000 /entries/Stone_Relation_Algebras.html Refining Authenticated Key Agreement with Strong Adversaries /entries/Key_Agreement_Strong_Adversaries.html Tue, 31 Jan 2017 00:00:00 +0000 /entries/Key_Agreement_Strong_Adversaries.html Bernoulli Numbers /entries/Bernoulli.html Tue, 24 Jan 2017 00:00:00 +0000 /entries/Bernoulli.html Bertrand's postulate /entries/Bertrands_Postulate.html Tue, 17 Jan 2017 00:00:00 +0000 /entries/Bertrands_Postulate.html Minimal Static Single Assignment Form /entries/Minimal_SSA.html Tue, 17 Jan 2017 00:00:00 +0000 /entries/Minimal_SSA.html The Transcendence of e /entries/E_Transcendental.html Thu, 12 Jan 2017 00:00:00 +0000 /entries/E_Transcendental.html Formal Network Models and Their Application to Firewall Policies /entries/UPF_Firewall.html Sun, 08 Jan 2017 00:00:00 +0000 /entries/UPF_Firewall.html Verification of a Diffie-Hellman Password-based Authentication Protocol by Extending the Inductive Method /entries/Password_Authentication_Protocol.html Tue, 03 Jan 2017 00:00:00 +0000 /entries/Password_Authentication_Protocol.html First-Order Logic According to Harrison /entries/FOL_Harrison.html Sun, 01 Jan 2017 00:00:00 +0000 /entries/FOL_Harrison.html Concurrent Refinement Algebra and Rely Quotients /entries/Concurrent_Ref_Alg.html Fri, 30 Dec 2016 00:00:00 +0000 /entries/Concurrent_Ref_Alg.html The Twelvefold Way /entries/Twelvefold_Way.html Thu, 29 Dec 2016 00:00:00 +0000 /entries/Twelvefold_Way.html Proof Strategy Language /entries/Proof_Strategy_Language.html Tue, 20 Dec 2016 00:00:00 +0000 /entries/Proof_Strategy_Language.html Paraconsistency /entries/Paraconsistency.html Wed, 07 Dec 2016 00:00:00 +0000 /entries/Paraconsistency.html COMPLX: A Verification Framework for Concurrent Imperative Programs /entries/Complx.html Tue, 29 Nov 2016 00:00:00 +0000 /entries/Complx.html Abstract Interpretation of Annotated Commands /entries/Abs_Int_ITP2012.html Wed, 23 Nov 2016 00:00:00 +0000 /entries/Abs_Int_ITP2012.html Separata: Isabelle tactics for Separation Algebra /entries/Separata.html Wed, 16 Nov 2016 00:00:00 +0000 /entries/Separata.html Formalization of Knuth–Bendix Orders for Lambda-Free Higher-Order Terms /entries/Lambda_Free_KBOs.html Sat, 12 Nov 2016 00:00:00 +0000 /entries/Lambda_Free_KBOs.html Formalization of Nested Multisets, Hereditary Multisets, and Syntactic Ordinals /entries/Nested_Multisets_Ordinals.html Sat, 12 Nov 2016 00:00:00 +0000 /entries/Nested_Multisets_Ordinals.html Expressiveness of Deep Learning /entries/Deep_Learning.html Thu, 10 Nov 2016 00:00:00 +0000 /entries/Deep_Learning.html Modal Logics for Nominal Transition Systems /entries/Modal_Logics_for_NTS.html Tue, 25 Oct 2016 00:00:00 +0000 /entries/Modal_Logics_for_NTS.html Stable Matching /entries/Stable_Matching.html Mon, 24 Oct 2016 00:00:00 +0000 /entries/Stable_Matching.html LOFT — Verified Migration of Linux Firewalls to SDN /entries/LOFT.html Fri, 21 Oct 2016 00:00:00 +0000 /entries/LOFT.html A formal model for the SPARCv8 ISA and a proof of non-interference for the LEON3 processor /entries/SPARCv8.html Wed, 19 Oct 2016 00:00:00 +0000 /entries/SPARCv8.html Source Coding Theorem /entries/Source_Coding_Theorem.html Wed, 19 Oct 2016 00:00:00 +0000 /entries/Source_Coding_Theorem.html The Factorization Algorithm of Berlekamp and Zassenhaus /entries/Berlekamp_Zassenhaus.html Fri, 14 Oct 2016 00:00:00 +0000 /entries/Berlekamp_Zassenhaus.html Intersecting Chords Theorem /entries/Chord_Segments.html Tue, 11 Oct 2016 00:00:00 +0000 /entries/Chord_Segments.html Lp spaces /entries/Lp.html Wed, 05 Oct 2016 00:00:00 +0000 /entries/Lp.html Fisher–Yates shuffle /entries/Fisher_Yates.html Fri, 30 Sep 2016 00:00:00 +0000 /entries/Fisher_Yates.html Allen's Interval Calculus /entries/Allen_Calculus.html Thu, 29 Sep 2016 00:00:00 +0000 /entries/Allen_Calculus.html Formalization of Recursive Path Orders for Lambda-Free Higher-Order Terms /entries/Lambda_Free_RPOs.html Fri, 23 Sep 2016 00:00:00 +0000 /entries/Lambda_Free_RPOs.html Iptables Semantics /entries/Iptables_Semantics.html Fri, 09 Sep 2016 00:00:00 +0000 /entries/Iptables_Semantics.html A Variant of the Superposition Calculus /entries/SuperCalc.html Tue, 06 Sep 2016 00:00:00 +0000 /entries/SuperCalc.html Stone Algebras /entries/Stone_Algebras.html Tue, 06 Sep 2016 00:00:00 +0000 /entries/Stone_Algebras.html Stirling's formula /entries/Stirling_Formula.html Thu, 01 Sep 2016 00:00:00 +0000 /entries/Stirling_Formula.html Routing /entries/Routing.html Wed, 31 Aug 2016 00:00:00 +0000 /entries/Routing.html Simple Firewall /entries/Simple_Firewall.html Wed, 24 Aug 2016 00:00:00 +0000 /entries/Simple_Firewall.html Infeasible Paths Elimination by Symbolic Execution Techniques: Proof of Correctness and Preservation of Paths /entries/InfPathElimination.html Thu, 18 Aug 2016 00:00:00 +0000 /entries/InfPathElimination.html Formalizing the Edmonds-Karp Algorithm /entries/EdmondsKarp_Maxflow.html Fri, 12 Aug 2016 00:00:00 +0000 /entries/EdmondsKarp_Maxflow.html The Imperative Refinement Framework /entries/Refine_Imperative_HOL.html Mon, 08 Aug 2016 00:00:00 +0000 /entries/Refine_Imperative_HOL.html Ptolemy's Theorem /entries/Ptolemys_Theorem.html Sun, 07 Aug 2016 00:00:00 +0000 /entries/Ptolemys_Theorem.html Surprise Paradox /entries/Surprise_Paradox.html Sun, 17 Jul 2016 00:00:00 +0000 /entries/Surprise_Paradox.html Pairing Heap /entries/Pairing_Heap.html Thu, 14 Jul 2016 00:00:00 +0000 /entries/Pairing_Heap.html A Framework for Verifying Depth-First Search Algorithms /entries/DFS_Framework.html Tue, 05 Jul 2016 00:00:00 +0000 /entries/DFS_Framework.html Chamber Complexes, Coxeter Systems, and Buildings /entries/Buildings.html Fri, 01 Jul 2016 00:00:00 +0000 /entries/Buildings.html The Resolution Calculus for First-Order Logic /entries/Resolution_FOL.html Thu, 30 Jun 2016 00:00:00 +0000 /entries/Resolution_FOL.html The Z Property /entries/Rewriting_Z.html Thu, 30 Jun 2016 00:00:00 +0000 /entries/Rewriting_Z.html Compositional Security-Preserving Refinement for Concurrent Imperative Programs /entries/Dependent_SIFUM_Refinement.html Tue, 28 Jun 2016 00:00:00 +0000 /entries/Dependent_SIFUM_Refinement.html IP Addresses /entries/IP_Addresses.html Tue, 28 Jun 2016 00:00:00 +0000 /entries/IP_Addresses.html Cardinality of Multisets /entries/Card_Multisets.html Sun, 26 Jun 2016 00:00:00 +0000 /entries/Card_Multisets.html Category Theory with Adjunctions and Limits /entries/Category3.html Sun, 26 Jun 2016 00:00:00 +0000 /entries/Category3.html A Dependent Security Type System for Concurrent Imperative Programs /entries/Dependent_SIFUM_Type_Systems.html Sat, 25 Jun 2016 00:00:00 +0000 /entries/Dependent_SIFUM_Type_Systems.html Catalan Numbers /entries/Catalan_Numbers.html Tue, 21 Jun 2016 00:00:00 +0000 /entries/Catalan_Numbers.html Program Construction and Verification Components Based on Kleene Algebra /entries/Algebraic_VCs.html Sat, 18 Jun 2016 00:00:00 +0000 /entries/Algebraic_VCs.html Conservation of CSP Noninterference Security under Concurrent Composition /entries/Noninterference_Concurrent_Composition.html Mon, 13 Jun 2016 00:00:00 +0000 /entries/Noninterference_Concurrent_Composition.html Finite Machine Word Library /entries/Word_Lib.html Thu, 09 Jun 2016 00:00:00 +0000 /entries/Word_Lib.html Tree Decomposition /entries/Tree_Decomposition.html Tue, 31 May 2016 00:00:00 +0000 /entries/Tree_Decomposition.html Cardinality of Equivalence Relations /entries/Card_Equiv_Relations.html Tue, 24 May 2016 00:00:00 +0000 /entries/Card_Equiv_Relations.html POSIX Lexing with Derivatives of Regular Expressions /entries/Posix-Lexing.html Tue, 24 May 2016 00:00:00 +0000 /entries/Posix-Lexing.html Perron-Frobenius Theorem for Spectral Radius Analysis /entries/Perron_Frobenius.html Fri, 20 May 2016 00:00:00 +0000 /entries/Perron_Frobenius.html The meta theory of the Incredible Proof Machine /entries/Incredible_Proof_Machine.html Fri, 20 May 2016 00:00:00 +0000 /entries/Incredible_Proof_Machine.html A Constructive Proof for FLP /entries/FLP.html Wed, 18 May 2016 00:00:00 +0000 /entries/FLP.html A Formal Proof of the Max-Flow Min-Cut Theorem for Countable Networks /entries/MFMC_Countable.html Mon, 09 May 2016 00:00:00 +0000 /entries/MFMC_Countable.html Randomised Social Choice Theory /entries/Randomised_Social_Choice.html Thu, 05 May 2016 00:00:00 +0000 /entries/Randomised_Social_Choice.html Spivey's Generalized Recurrence for Bell Numbers /entries/Bell_Numbers_Spivey.html Wed, 04 May 2016 00:00:00 +0000 /entries/Bell_Numbers_Spivey.html The Incompatibility of SD-Efficiency and SD-Strategy-Proofness /entries/SDS_Impossibility.html Wed, 04 May 2016 00:00:00 +0000 /entries/SDS_Impossibility.html Gröbner Bases Theory /entries/Groebner_Bases.html Mon, 02 May 2016 00:00:00 +0000 /entries/Groebner_Bases.html No Faster-Than-Light Observers /entries/No_FTL_observers.html Thu, 28 Apr 2016 00:00:00 +0000 /entries/No_FTL_observers.html A formalisation of the Cocke-Younger-Kasami algorithm /entries/CYK.html Wed, 27 Apr 2016 00:00:00 +0000 /entries/CYK.html Algorithms for Reduced Ordered Binary Decision Diagrams /entries/ROBDD.html Wed, 27 Apr 2016 00:00:00 +0000 /entries/ROBDD.html Conservation of CSP Noninterference Security under Sequential Composition /entries/Noninterference_Sequential_Composition.html Tue, 26 Apr 2016 00:00:00 +0000 /entries/Noninterference_Sequential_Composition.html Kleene Algebras with Domain /entries/KAD.html Tue, 12 Apr 2016 00:00:00 +0000 /entries/KAD.html Propositional Resolution and Prime Implicates Generation /entries/PropResPI.html Fri, 11 Mar 2016 00:00:00 +0000 /entries/PropResPI.html The Cartan Fixed Point Theorems /entries/Cartan_FP.html Tue, 08 Mar 2016 00:00:00 +0000 /entries/Cartan_FP.html Timed Automata /entries/Timed_Automata.html Tue, 08 Mar 2016 00:00:00 +0000 /entries/Timed_Automata.html Linear Temporal Logic /entries/LTL.html Tue, 01 Mar 2016 00:00:00 +0000 /entries/LTL.html Analysis of List Update Algorithms /entries/List_Update.html Wed, 17 Feb 2016 00:00:00 +0000 /entries/List_Update.html Verified Construction of Static Single Assignment Form /entries/Formal_SSA.html Fri, 05 Feb 2016 00:00:00 +0000 /entries/Formal_SSA.html Polynomial Factorization /entries/Polynomial_Factorization.html Fri, 29 Jan 2016 00:00:00 +0000 /entries/Polynomial_Factorization.html Polynomial Interpolation /entries/Polynomial_Interpolation.html Fri, 29 Jan 2016 00:00:00 +0000 /entries/Polynomial_Interpolation.html Knot Theory /entries/Knot_Theory.html Wed, 20 Jan 2016 00:00:00 +0000 /entries/Knot_Theory.html Tensor Product of Matrices /entries/Matrix_Tensor.html Mon, 18 Jan 2016 00:00:00 +0000 /entries/Matrix_Tensor.html Cardinality of Number Partitions /entries/Card_Number_Partitions.html Thu, 14 Jan 2016 00:00:00 +0000 /entries/Card_Number_Partitions.html Basic Geometric Properties of Triangles /entries/Triangle.html Mon, 28 Dec 2015 00:00:00 +0000 /entries/Triangle.html Descartes' Rule of Signs /entries/Descartes_Sign_Rule.html Mon, 28 Dec 2015 00:00:00 +0000 /entries/Descartes_Sign_Rule.html Liouville numbers /entries/Liouville_Numbers.html Mon, 28 Dec 2015 00:00:00 +0000 /entries/Liouville_Numbers.html The Divergence of the Prime Harmonic Series /entries/Prime_Harmonic_Series.html Mon, 28 Dec 2015 00:00:00 +0000 /entries/Prime_Harmonic_Series.html Algebraic Numbers in Isabelle/HOL /entries/Algebraic_Numbers.html Tue, 22 Dec 2015 00:00:00 +0000 /entries/Algebraic_Numbers.html Applicative Lifting /entries/Applicative_Lifting.html Tue, 22 Dec 2015 00:00:00 +0000 /entries/Applicative_Lifting.html The Stern-Brocot Tree /entries/Stern_Brocot.html Tue, 22 Dec 2015 00:00:00 +0000 /entries/Stern_Brocot.html Cardinality of Set Partitions /entries/Card_Partitions.html Sat, 12 Dec 2015 00:00:00 +0000 /entries/Card_Partitions.html Latin Square /entries/Latin_Square.html Wed, 02 Dec 2015 00:00:00 +0000 /entries/Latin_Square.html Ergodic Theory /entries/Ergodic_Theory.html Tue, 01 Dec 2015 00:00:00 +0000 /entries/Ergodic_Theory.html Euler's Partition Theorem /entries/Euler_Partition.html Thu, 19 Nov 2015 00:00:00 +0000 /entries/Euler_Partition.html The Tortoise and Hare Algorithm /entries/TortoiseHare.html Wed, 18 Nov 2015 00:00:00 +0000 /entries/TortoiseHare.html Planarity Certificates /entries/Planarity_Certificates.html Wed, 11 Nov 2015 00:00:00 +0000 /entries/Planarity_Certificates.html Positional Determinacy of Parity Games /entries/Parity_Game.html Mon, 02 Nov 2015 00:00:00 +0000 /entries/Parity_Game.html A Meta-Model for the Isabelle API /entries/Isabelle_Meta_Model.html Wed, 16 Sep 2015 00:00:00 +0000 /entries/Isabelle_Meta_Model.html Converting Linear Temporal Logic to Deterministic (Generalized) Rabin Automata /entries/LTL_to_DRA.html Fri, 04 Sep 2015 00:00:00 +0000 /entries/LTL_to_DRA.html Matrices, Jordan Normal Forms, and Spectral Radius Theory /entries/Jordan_Normal_Form.html Fri, 21 Aug 2015 00:00:00 +0000 /entries/Jordan_Normal_Form.html Decreasing Diagrams II /entries/Decreasing-Diagrams-II.html Thu, 20 Aug 2015 00:00:00 +0000 /entries/Decreasing-Diagrams-II.html The Inductive Unwinding Theorem for CSP Noninterference Security /entries/Noninterference_Inductive_Unwinding.html Tue, 18 Aug 2015 00:00:00 +0000 /entries/Noninterference_Inductive_Unwinding.html Representations of Finite Groups /entries/Rep_Fin_Groups.html Wed, 12 Aug 2015 00:00:00 +0000 /entries/Rep_Fin_Groups.html Analysing and Comparing Encodability Criteria for Process Calculi /entries/Encodability_Process_Calculi.html Mon, 10 Aug 2015 00:00:00 +0000 /entries/Encodability_Process_Calculi.html Generating Cases from Labeled Subgoals /entries/Case_Labeling.html Tue, 21 Jul 2015 00:00:00 +0000 /entries/Case_Labeling.html Landau Symbols /entries/Landau_Symbols.html Tue, 14 Jul 2015 00:00:00 +0000 /entries/Landau_Symbols.html The Akra-Bazzi theorem and the Master theorem /entries/Akra_Bazzi.html Tue, 14 Jul 2015 00:00:00 +0000 /entries/Akra_Bazzi.html Hermite Normal Form /entries/Hermite.html Tue, 07 Jul 2015 00:00:00 +0000 /entries/Hermite.html Derangements Formula /entries/Derangements.html Sat, 27 Jun 2015 00:00:00 +0000 /entries/Derangements.html Binary Multirelations /entries/Multirelations.html Thu, 11 Jun 2015 00:00:00 +0000 /entries/Multirelations.html Reasoning about Lists via List Interleaving /entries/List_Interleaving.html Thu, 11 Jun 2015 00:00:00 +0000 /entries/List_Interleaving.html The Generic Unwinding Theorem for CSP Noninterference Security /entries/Noninterference_Generic_Unwinding.html Thu, 11 Jun 2015 00:00:00 +0000 /entries/Noninterference_Generic_Unwinding.html The Ipurge Unwinding Theorem for CSP Noninterference Security /entries/Noninterference_Ipurge_Unwinding.html Thu, 11 Jun 2015 00:00:00 +0000 /entries/Noninterference_Ipurge_Unwinding.html Parameterized Dynamic Tables /entries/Dynamic_Tables.html Sun, 07 Jun 2015 00:00:00 +0000 /entries/Dynamic_Tables.html Derivatives of Logical Formulas /entries/Formula_Derivatives.html Thu, 28 May 2015 00:00:00 +0000 /entries/Formula_Derivatives.html A Zoo of Probabilistic Systems /entries/Probabilistic_System_Zoo.html Wed, 27 May 2015 00:00:00 +0000 /entries/Probabilistic_System_Zoo.html VCG - Combinatorial Vickrey-Clarke-Groves Auctions /entries/Vickrey_Clarke_Groves.html Thu, 30 Apr 2015 00:00:00 +0000 /entries/Vickrey_Clarke_Groves.html Residuated Lattices /entries/Residuated_Lattices.html Wed, 15 Apr 2015 00:00:00 +0000 /entries/Residuated_Lattices.html Concurrent IMP /entries/ConcurrentIMP.html Mon, 13 Apr 2015 00:00:00 +0000 /entries/ConcurrentIMP.html Relaxing Safely: Verified On-the-Fly Garbage Collection for x86-TSO /entries/ConcurrentGC.html Mon, 13 Apr 2015 00:00:00 +0000 /entries/ConcurrentGC.html Trie /entries/Trie.html Mon, 30 Mar 2015 00:00:00 +0000 /entries/Trie.html Consensus Refined /entries/Consensus_Refined.html Wed, 18 Mar 2015 00:00:00 +0000 /entries/Consensus_Refined.html Deriving class instances for datatypes /entries/Deriving.html Wed, 11 Mar 2015 00:00:00 +0000 /entries/Deriving.html The Safety of Call Arity /entries/Call_Arity.html Fri, 20 Feb 2015 00:00:00 +0000 /entries/Call_Arity.html Echelon Form /entries/Echelon_Form.html Thu, 12 Feb 2015 00:00:00 +0000 /entries/Echelon_Form.html QR Decomposition /entries/QR_Decomposition.html Thu, 12 Feb 2015 00:00:00 +0000 /entries/QR_Decomposition.html Finite Automata in Hereditarily Finite Set Theory /entries/Finite_Automata_HF.html Thu, 05 Feb 2015 00:00:00 +0000 /entries/Finite_Automata_HF.html Verification of the UpDown Scheme /entries/UpDown_Scheme.html Wed, 28 Jan 2015 00:00:00 +0000 /entries/UpDown_Scheme.html The Unified Policy Framework (UPF) /entries/UPF.html Fri, 28 Nov 2014 00:00:00 +0000 /entries/UPF.html Loop freedom of the (untimed) AODV routing protocol /entries/AODV.html Thu, 23 Oct 2014 00:00:00 +0000 /entries/AODV.html Lifting Definition Option /entries/Lifting_Definition_Option.html Mon, 13 Oct 2014 00:00:00 +0000 /entries/Lifting_Definition_Option.html Stream Fusion in HOL with Code Generation /entries/Stream_Fusion_Code.html Fri, 10 Oct 2014 00:00:00 +0000 /entries/Stream_Fusion_Code.html A Verified Compiler for Probability Density Functions /entries/Density_Compiler.html Thu, 09 Oct 2014 00:00:00 +0000 /entries/Density_Compiler.html Formalization of Refinement Calculus for Reactive Systems /entries/RefinementReactive.html Wed, 08 Oct 2014 00:00:00 +0000 /entries/RefinementReactive.html Certification Monads /entries/Certification_Monads.html Fri, 03 Oct 2014 00:00:00 +0000 /entries/Certification_Monads.html XML /entries/XML.html Fri, 03 Oct 2014 00:00:00 +0000 /entries/XML.html Imperative Insertion Sort /entries/Imperative_Insertion_Sort.html Thu, 25 Sep 2014 00:00:00 +0000 /entries/Imperative_Insertion_Sort.html The Sturm-Tarski Theorem /entries/Sturm_Tarski.html Fri, 19 Sep 2014 00:00:00 +0000 /entries/Sturm_Tarski.html The Cayley-Hamilton Theorem /entries/Cayley_Hamilton.html Mon, 15 Sep 2014 00:00:00 +0000 /entries/Cayley_Hamilton.html The Jordan-Hölder Theorem /entries/Jordan_Hoelder.html Tue, 09 Sep 2014 00:00:00 +0000 /entries/Jordan_Hoelder.html Priority Queues Based on Braun Trees /entries/Priority_Queue_Braun.html Thu, 04 Sep 2014 00:00:00 +0000 /entries/Priority_Queue_Braun.html Gauss-Jordan Algorithm and Its Applications /entries/Gauss_Jordan.html Wed, 03 Sep 2014 00:00:00 +0000 /entries/Gauss_Jordan.html Real-Valued Special Functions: Upper and Lower Bounds /entries/Special_Function_Bounds.html Fri, 29 Aug 2014 00:00:00 +0000 /entries/Special_Function_Bounds.html Vector Spaces /entries/VectorSpace.html Fri, 29 Aug 2014 00:00:00 +0000 /entries/VectorSpace.html Skew Heap /entries/Skew_Heap.html Wed, 13 Aug 2014 00:00:00 +0000 /entries/Skew_Heap.html Splay Tree /entries/Splay_Tree.html Tue, 12 Aug 2014 00:00:00 +0000 /entries/Splay_Tree.html Haskell's Show Class in Isabelle/HOL /entries/Show.html Tue, 29 Jul 2014 00:00:00 +0000 /entries/Show.html Formal Specification of a Generic Separation Kernel /entries/CISC-Kernel.html Fri, 18 Jul 2014 00:00:00 +0000 /entries/CISC-Kernel.html pGCL for Isabelle /entries/pGCL.html Sun, 13 Jul 2014 00:00:00 +0000 /entries/pGCL.html Amortized Complexity Verified /entries/Amortized_Complexity.html Mon, 07 Jul 2014 00:00:00 +0000 /entries/Amortized_Complexity.html Network Security Policy Verification /entries/Network_Security_Policy_Verification.html Fri, 04 Jul 2014 00:00:00 +0000 /entries/Network_Security_Policy_Verification.html Pop-Refinement /entries/Pop_Refinement.html Thu, 03 Jul 2014 00:00:00 +0000 /entries/Pop_Refinement.html Decision Procedures for MSO on Words Based on Derivatives of Regular Expressions /entries/MSO_Regex_Equivalence.html Thu, 12 Jun 2014 00:00:00 +0000 /entries/MSO_Regex_Equivalence.html Boolean Expression Checkers /entries/Boolean_Expression_Checkers.html Sun, 08 Jun 2014 00:00:00 +0000 /entries/Boolean_Expression_Checkers.html A Fully Verified Executable LTL Model Checker /entries/CAVA_LTL_Modelchecker.html Wed, 28 May 2014 00:00:00 +0000 /entries/CAVA_LTL_Modelchecker.html Converting Linear-Time Temporal Logic to Generalized Büchi Automata /entries/LTL_to_GBA.html Wed, 28 May 2014 00:00:00 +0000 /entries/LTL_to_GBA.html Promela Formalization /entries/Promela.html Wed, 28 May 2014 00:00:00 +0000 /entries/Promela.html The CAVA Automata Library /entries/CAVA_Automata.html Wed, 28 May 2014 00:00:00 +0000 /entries/CAVA_Automata.html Verified Efficient Implementation of Gabow's Strongly Connected Components Algorithm /entries/Gabow_SCC.html Wed, 28 May 2014 00:00:00 +0000 /entries/Gabow_SCC.html Noninterference Security in Communicating Sequential Processes /entries/Noninterference_CSP.html Fri, 23 May 2014 00:00:00 +0000 /entries/Noninterference_CSP.html Transitive closure according to Roy-Floyd-Warshall /entries/Roy_Floyd_Warshall.html Fri, 23 May 2014 00:00:00 +0000 /entries/Roy_Floyd_Warshall.html Regular Algebras /entries/Regular_Algebras.html Wed, 21 May 2014 00:00:00 +0000 /entries/Regular_Algebras.html Formalisation and Analysis of Component Dependencies /entries/ComponentDependencies.html Mon, 28 Apr 2014 00:00:00 +0000 /entries/ComponentDependencies.html A Formalization of Assumptions and Guarantees for Compositional Noninterference /entries/SIFUM_Type_Systems.html Wed, 23 Apr 2014 00:00:00 +0000 /entries/SIFUM_Type_Systems.html A Formalization of Declassification with WHAT-and-WHERE-Security /entries/WHATandWHERE_Security.html Wed, 23 Apr 2014 00:00:00 +0000 /entries/WHATandWHERE_Security.html A Formalization of Strong Security /entries/Strong_Security.html Wed, 23 Apr 2014 00:00:00 +0000 /entries/Strong_Security.html Bounded-Deducibility Security /entries/Bounded_Deducibility_Security.html Tue, 22 Apr 2014 00:00:00 +0000 /entries/Bounded_Deducibility_Security.html A shallow embedding of HyperCTL* /entries/HyperCTL.html Wed, 16 Apr 2014 00:00:00 +0000 /entries/HyperCTL.html Abstract Completeness /entries/Abstract_Completeness.html Wed, 16 Apr 2014 00:00:00 +0000 /entries/Abstract_Completeness.html Discrete Summation /entries/Discrete_Summation.html Sun, 13 Apr 2014 00:00:00 +0000 /entries/Discrete_Summation.html Syntax and semantics of a GPU kernel programming language /entries/GPU_Kernel_PL.html Thu, 03 Apr 2014 00:00:00 +0000 /entries/GPU_Kernel_PL.html Probabilistic Noninterference /entries/Probabilistic_Noninterference.html Tue, 11 Mar 2014 00:00:00 +0000 /entries/Probabilistic_Noninterference.html Mechanization of the Algebra for Wireless Networks (AWN) /entries/AWN.html Sat, 08 Mar 2014 00:00:00 +0000 /entries/AWN.html Mutually Recursive Partial Functions /entries/Partial_Function_MR.html Tue, 18 Feb 2014 00:00:00 +0000 /entries/Partial_Function_MR.html Properties of Random Graphs -- Subgraph Containment /entries/Random_Graph_Subgraph_Threshold.html Thu, 13 Feb 2014 00:00:00 +0000 /entries/Random_Graph_Subgraph_Threshold.html Verification of Selection and Heap Sort Using Locales /entries/Selection_Heap_Sort.html Tue, 11 Feb 2014 00:00:00 +0000 /entries/Selection_Heap_Sort.html Affine Arithmetic /entries/Affine_Arithmetic.html Fri, 07 Feb 2014 00:00:00 +0000 /entries/Affine_Arithmetic.html Implementing field extensions of the form Q[sqrt(b)] /entries/Real_Impl.html Thu, 06 Feb 2014 00:00:00 +0000 /entries/Real_Impl.html Unified Decision Procedures for Regular Expression Equivalence /entries/Regex_Equivalence.html Thu, 30 Jan 2014 00:00:00 +0000 /entries/Regex_Equivalence.html Secondary Sylow Theorems /entries/Secondary_Sylow.html Tue, 28 Jan 2014 00:00:00 +0000 /entries/Secondary_Sylow.html Relation Algebra /entries/Relation_Algebra.html Sat, 25 Jan 2014 00:00:00 +0000 /entries/Relation_Algebra.html Kleene Algebra with Tests and Demonic Refinement Algebras /entries/KAT_and_DRA.html Thu, 23 Jan 2014 00:00:00 +0000 /entries/KAT_and_DRA.html Featherweight OCL: A Proposal for a Machine-Checked Formal Semantics for OCL 2.5 /entries/Featherweight_OCL.html Thu, 16 Jan 2014 00:00:00 +0000 /entries/Featherweight_OCL.html Compositional Properties of Crypto-Based Components /entries/CryptoBasedCompositionalProperties.html Sat, 11 Jan 2014 00:00:00 +0000 /entries/CryptoBasedCompositionalProperties.html Sturm's Theorem /entries/Sturm_Sequences.html Sat, 11 Jan 2014 00:00:00 +0000 /entries/Sturm_Sequences.html A General Method for the Proof of Theorems on Tail-recursive Functions /entries/Tail_Recursive_Functions.html Sun, 01 Dec 2013 00:00:00 +0000 /entries/Tail_Recursive_Functions.html Gödel's Incompleteness Theorems /entries/Incompleteness.html Sun, 17 Nov 2013 00:00:00 +0000 /entries/Incompleteness.html The Hereditarily Finite Sets /entries/HereditarilyFinite.html Sun, 17 Nov 2013 00:00:00 +0000 /entries/HereditarilyFinite.html A Codatatype of Formal Languages /entries/Coinductive_Languages.html Fri, 15 Nov 2013 00:00:00 +0000 /entries/Coinductive_Languages.html Stream Processing Components: Isabelle/HOL Formalisation and Case Studies /entries/FocusStreamsCaseStudies.html Thu, 14 Nov 2013 00:00:00 +0000 /entries/FocusStreamsCaseStudies.html Gödel's God in Isabelle/HOL /entries/GoedelGod.html Tue, 12 Nov 2013 00:00:00 +0000 /entries/GoedelGod.html Decreasing Diagrams /entries/Decreasing-Diagrams.html Fri, 01 Nov 2013 00:00:00 +0000 /entries/Decreasing-Diagrams.html Automatic Data Refinement /entries/Automatic_Refinement.html Wed, 02 Oct 2013 00:00:00 +0000 /entries/Automatic_Refinement.html Native Word /entries/Native_Word.html Tue, 17 Sep 2013 00:00:00 +0000 /entries/Native_Word.html A Formal Model of IEEE Floating Point Arithmetic /entries/IEEE_Floating_Point.html Sat, 27 Jul 2013 00:00:00 +0000 /entries/IEEE_Floating_Point.html Lehmer's Theorem /entries/Lehmer.html Mon, 22 Jul 2013 00:00:00 +0000 /entries/Lehmer.html Pratt's Primality Certificates /entries/Pratt_Certificate.html Mon, 22 Jul 2013 00:00:00 +0000 /entries/Pratt_Certificate.html The Königsberg Bridge Problem and the Friendship Theorem /entries/Koenigsberg_Friendship.html Fri, 19 Jul 2013 00:00:00 +0000 /entries/Koenigsberg_Friendship.html Sound and Complete Sort Encodings for First-Order Logic /entries/Sort_Encodings.html Thu, 27 Jun 2013 00:00:00 +0000 /entries/Sort_Encodings.html An Axiomatic Characterization of the Single-Source Shortest Path Problem /entries/ShortestPath.html Wed, 22 May 2013 00:00:00 +0000 /entries/ShortestPath.html Graph Theory /entries/Graph_Theory.html Sun, 28 Apr 2013 00:00:00 +0000 /entries/Graph_Theory.html Light-weight Containers /entries/Containers.html Mon, 15 Apr 2013 00:00:00 +0000 /entries/Containers.html Nominal 2 /entries/Nominal2.html Thu, 21 Feb 2013 00:00:00 +0000 /entries/Nominal2.html The Correctness of Launchbury's Natural Semantics for Lazy Evaluation /entries/Launchbury.html Thu, 31 Jan 2013 00:00:00 +0000 /entries/Launchbury.html Ribbon Proofs /entries/Ribbon_Proofs.html Sat, 19 Jan 2013 00:00:00 +0000 /entries/Ribbon_Proofs.html Rank-Nullity Theorem in Linear Algebra /entries/Rank_Nullity_Theorem.html Wed, 16 Jan 2013 00:00:00 +0000 /entries/Rank_Nullity_Theorem.html Kleene Algebra /entries/Kleene_Algebra.html Tue, 15 Jan 2013 00:00:00 +0000 /entries/Kleene_Algebra.html Computing N-th Roots using the Babylonian Method /entries/Sqrt_Babylonian.html Thu, 03 Jan 2013 00:00:00 +0000 /entries/Sqrt_Babylonian.html A Separation Logic Framework for Imperative HOL /entries/Separation_Logic_Imperative_HOL.html Wed, 14 Nov 2012 00:00:00 +0000 /entries/Separation_Logic_Imperative_HOL.html Open Induction /entries/Open_Induction.html Fri, 02 Nov 2012 00:00:00 +0000 /entries/Open_Induction.html The independence of Tarski's Euclidean axiom /entries/Tarskis_Geometry.html Tue, 30 Oct 2012 00:00:00 +0000 /entries/Tarskis_Geometry.html Bondy's Theorem /entries/Bondy.html Sat, 27 Oct 2012 00:00:00 +0000 /entries/Bondy.html Possibilistic Noninterference /entries/Possibilistic_Noninterference.html Mon, 10 Sep 2012 00:00:00 +0000 /entries/Possibilistic_Noninterference.html Generating linear orders for datatypes /entries/Datatype_Order_Generator.html Tue, 07 Aug 2012 00:00:00 +0000 /entries/Datatype_Order_Generator.html Proving the Impossibility of Trisecting an Angle and Doubling the Cube /entries/Impossible_Geometry.html Sun, 05 Aug 2012 00:00:00 +0000 /entries/Impossible_Geometry.html Verifying Fault-Tolerant Distributed Algorithms in the Heard-Of Model /entries/Heard_Of.html Fri, 27 Jul 2012 00:00:00 +0000 /entries/Heard_Of.html Logical Relations for PCF /entries/PCF.html Sun, 01 Jul 2012 00:00:00 +0000 /entries/PCF.html Type Constructor Classes and Monad Transformers /entries/Tycon.html Tue, 26 Jun 2012 00:00:00 +0000 /entries/Tycon.html CCS in nominal logic /entries/CCS.html Tue, 29 May 2012 00:00:00 +0000 /entries/CCS.html Psi-calculi in Isabelle /entries/Psi_Calculi.html Tue, 29 May 2012 00:00:00 +0000 /entries/Psi_Calculi.html The pi-calculus in nominal logic /entries/Pi_Calculus.html Tue, 29 May 2012 00:00:00 +0000 /entries/Pi_Calculus.html Isabelle/Circus /entries/Circus.html Sun, 27 May 2012 00:00:00 +0000 /entries/Circus.html Separation Algebra /entries/Separation_Algebra.html Fri, 11 May 2012 00:00:00 +0000 /entries/Separation_Algebra.html Stuttering Equivalence /entries/Stuttering_Equivalence.html Mon, 07 May 2012 00:00:00 +0000 /entries/Stuttering_Equivalence.html Inductive Study of Confidentiality /entries/Inductive_Confidentiality.html Wed, 02 May 2012 00:00:00 +0000 /entries/Inductive_Confidentiality.html Ordinary Differential Equations /entries/Ordinary_Differential_Equations.html Thu, 26 Apr 2012 00:00:00 +0000 /entries/Ordinary_Differential_Equations.html Well-Quasi-Orders /entries/Well_Quasi_Orders.html Fri, 13 Apr 2012 00:00:00 +0000 /entries/Well_Quasi_Orders.html Abortable Linearizable Modules /entries/Abortable_Linearizable_Modules.html Thu, 01 Mar 2012 00:00:00 +0000 /entries/Abortable_Linearizable_Modules.html Executable Transitive Closures /entries/Transitive-Closure-II.html Wed, 29 Feb 2012 00:00:00 +0000 /entries/Transitive-Closure-II.html A Probabilistic Proof of the Girth-Chromatic Number Theorem /entries/Girth_Chromatic.html Mon, 06 Feb 2012 00:00:00 +0000 /entries/Girth_Chromatic.html Dijkstra's Shortest Path Algorithm /entries/Dijkstra_Shortest_Path.html Mon, 30 Jan 2012 00:00:00 +0000 /entries/Dijkstra_Shortest_Path.html Refinement for Monadic Programs /entries/Refine_Monadic.html Mon, 30 Jan 2012 00:00:00 +0000 /entries/Refine_Monadic.html Markov Models /entries/Markov_Models.html Tue, 03 Jan 2012 00:00:00 +0000 /entries/Markov_Models.html A Definitional Encoding of TLA* in Isabelle/HOL /entries/TLA.html Sat, 19 Nov 2011 00:00:00 +0000 /entries/TLA.html Efficient Mergesort /entries/Efficient-Mergesort.html Wed, 09 Nov 2011 00:00:00 +0000 /entries/Efficient-Mergesort.html Algebra of Monotonic Boolean Transformers /entries/MonoBoolTranAlgebra.html Thu, 22 Sep 2011 00:00:00 +0000 /entries/MonoBoolTranAlgebra.html Lattice Properties /entries/LatticeProperties.html Thu, 22 Sep 2011 00:00:00 +0000 /entries/LatticeProperties.html Pseudo Hoops /entries/PseudoHoops.html Thu, 22 Sep 2011 00:00:00 +0000 /entries/PseudoHoops.html The Myhill-Nerode Theorem Based on Regular Expressions /entries/Myhill-Nerode.html Fri, 26 Aug 2011 00:00:00 +0000 /entries/Myhill-Nerode.html Gauss-Jordan Elimination for Matrices Represented as Functions /entries/Gauss-Jordan-Elim-Fun.html Fri, 19 Aug 2011 00:00:00 +0000 /entries/Gauss-Jordan-Elim-Fun.html Maximum Cardinality Matching /entries/Max-Card-Matching.html Thu, 21 Jul 2011 00:00:00 +0000 /entries/Max-Card-Matching.html Knowledge-based programs /entries/KBPs.html Tue, 17 May 2011 00:00:00 +0000 /entries/KBPs.html The General Triangle Is Unique /entries/General-Triangle.html Fri, 01 Apr 2011 00:00:00 +0000 /entries/General-Triangle.html Executable Transitive Closures of Finite Relations /entries/Transitive-Closure.html Mon, 14 Mar 2011 00:00:00 +0000 /entries/Transitive-Closure.html AutoFocus Stream Processing for Single-Clocking and Multi-Clocking Semantics /entries/AutoFocus-Stream.html Wed, 23 Feb 2011 00:00:00 +0000 /entries/AutoFocus-Stream.html Infinite Lists /entries/List-Infinite.html Wed, 23 Feb 2011 00:00:00 +0000 /entries/List-Infinite.html Interval Temporal Logic on Natural Numbers /entries/Nat-Interval-Logic.html Wed, 23 Feb 2011 00:00:00 +0000 /entries/Nat-Interval-Logic.html Lightweight Java /entries/LightweightJava.html Mon, 07 Feb 2011 00:00:00 +0000 /entries/LightweightJava.html RIPEMD-160 /entries/RIPEMD-160-SPARK.html Mon, 10 Jan 2011 00:00:00 +0000 /entries/RIPEMD-160-SPARK.html Lower Semicontinuous Functions /entries/Lower_Semicontinuous.html Sat, 08 Jan 2011 00:00:00 +0000 /entries/Lower_Semicontinuous.html Hall's Marriage Theorem /entries/Marriage.html Fri, 17 Dec 2010 00:00:00 +0000 /entries/Marriage.html Shivers' Control Flow Analysis /entries/Shivers-CFA.html Tue, 16 Nov 2010 00:00:00 +0000 /entries/Shivers-CFA.html Binomial Heaps and Skew Binomial Heaps /entries/Binomial-Heaps.html Thu, 28 Oct 2010 00:00:00 +0000 /entries/Binomial-Heaps.html Finger Trees /entries/Finger-Trees.html Thu, 28 Oct 2010 00:00:00 +0000 /entries/Finger-Trees.html Functional Binomial Queues /entries/Binomial-Queues.html Thu, 28 Oct 2010 00:00:00 +0000 /entries/Binomial-Queues.html Strong Normalization of Moggis's Computational Metalanguage /entries/Lam-ml-Normalization.html Sun, 29 Aug 2010 00:00:00 +0000 /entries/Lam-ml-Normalization.html Executable Multivariate Polynomials /entries/Polynomials.html Tue, 10 Aug 2010 00:00:00 +0000 /entries/Polynomials.html Formalizing Statecharts using Hierarchical Automata /entries/Statecharts.html Sun, 08 Aug 2010 00:00:00 +0000 /entries/Statecharts.html Free Groups /entries/Free-Groups.html Thu, 24 Jun 2010 00:00:00 +0000 /entries/Free-Groups.html Category Theory /entries/Category2.html Sun, 20 Jun 2010 00:00:00 +0000 /entries/Category2.html Executable Matrix Operations on Matrices of Arbitrary Dimensions /entries/Matrix.html Thu, 17 Jun 2010 00:00:00 +0000 /entries/Matrix.html Abstract Rewriting /entries/Abstract-Rewriting.html Mon, 14 Jun 2010 00:00:00 +0000 /entries/Abstract-Rewriting.html Semantics and Data Refinement of Invariant Based Programs /entries/DataRefinementIBP.html Fri, 28 May 2010 00:00:00 +0000 /entries/DataRefinementIBP.html Verification of the Deutsch-Schorr-Waite Graph Marking Algorithm using Data Refinement /entries/GraphMarkingIBP.html Fri, 28 May 2010 00:00:00 +0000 /entries/GraphMarkingIBP.html A Complete Proof of the Robbins Conjecture /entries/Robbins-Conjecture.html Sat, 22 May 2010 00:00:00 +0000 /entries/Robbins-Conjecture.html Regular Sets and Expressions /entries/Regular-Sets.html Wed, 12 May 2010 00:00:00 +0000 /entries/Regular-Sets.html Locally Nameless Sigma Calculus /entries/Locally-Nameless-Sigma.html Fri, 30 Apr 2010 00:00:00 +0000 /entries/Locally-Nameless-Sigma.html Free Boolean Algebra /entries/Free-Boolean-Algebra.html Mon, 29 Mar 2010 00:00:00 +0000 /entries/Free-Boolean-Algebra.html Information Flow Noninterference via Slicing /entries/InformationFlowSlicing.html Tue, 23 Mar 2010 00:00:00 +0000 /entries/InformationFlowSlicing.html Inter-Procedural Information Flow Noninterference via Slicing /entries/InformationFlowSlicing_Inter.html Tue, 23 Mar 2010 00:00:00 +0000 /entries/InformationFlowSlicing_Inter.html List Index /entries/List-Index.html Sat, 20 Feb 2010 00:00:00 +0000 /entries/List-Index.html Coinductive /entries/Coinductive.html Fri, 12 Feb 2010 00:00:00 +0000 /entries/Coinductive.html A Fast SAT Solver for Isabelle in Standard ML /entries/DPT-SAT-Solver.html Wed, 09 Dec 2009 00:00:00 +0000 /entries/DPT-SAT-Solver.html Formalizing the Logic-Automaton Connection /entries/Presburger-Automata.html Thu, 03 Dec 2009 00:00:00 +0000 /entries/Presburger-Automata.html Collections Framework /entries/Collections.html Wed, 25 Nov 2009 00:00:00 +0000 /entries/Collections.html Tree Automata /entries/Tree-Automata.html Wed, 25 Nov 2009 00:00:00 +0000 /entries/Tree-Automata.html Perfect Number Theorem /entries/Perfect-Number-Thm.html Sun, 22 Nov 2009 00:00:00 +0000 /entries/Perfect-Number-Thm.html Backing up Slicing: Verifying the Interprocedural Two-Phase Horwitz-Reps-Binkley Slicer /entries/HRB-Slicing.html Fri, 13 Nov 2009 00:00:00 +0000 /entries/HRB-Slicing.html The Worker/Wrapper Transformation /entries/WorkerWrapper.html Fri, 30 Oct 2009 00:00:00 +0000 /entries/WorkerWrapper.html Ordinals and Cardinals /entries/Ordinals_and_Cardinals.html Tue, 01 Sep 2009 00:00:00 +0000 /entries/Ordinals_and_Cardinals.html Invertibility in Sequent Calculi /entries/SequentInvertibility.html Fri, 28 Aug 2009 00:00:00 +0000 /entries/SequentInvertibility.html An Example of a Cofinitary Group in Isabelle/HOL /entries/CofGroups.html Tue, 04 Aug 2009 00:00:00 +0000 /entries/CofGroups.html Code Generation for Functions as Data /entries/FinFun.html Wed, 06 May 2009 00:00:00 +0000 /entries/FinFun.html Stream Fusion /entries/Stream-Fusion.html Wed, 29 Apr 2009 00:00:00 +0000 /entries/Stream-Fusion.html A Bytecode Logic for JML and Types /entries/BytecodeLogicJmlTypes.html Fri, 12 Dec 2008 00:00:00 +0000 /entries/BytecodeLogicJmlTypes.html Secure information flow and program logics /entries/SIFPL.html Mon, 10 Nov 2008 00:00:00 +0000 /entries/SIFPL.html Some classical results in Social Choice Theory /entries/SenSocialChoice.html Sun, 09 Nov 2008 00:00:00 +0000 /entries/SenSocialChoice.html Fun With Tilings /entries/FunWithTilings.html Fri, 07 Nov 2008 00:00:00 +0000 /entries/FunWithTilings.html The Textbook Proof of Huffman's Algorithm /entries/Huffman.html Wed, 15 Oct 2008 00:00:00 +0000 /entries/Huffman.html Towards Certified Slicing /entries/Slicing.html Tue, 16 Sep 2008 00:00:00 +0000 /entries/Slicing.html A Correctness Proof for the Volpano/Smith Security Typing System /entries/VolpanoSmith.html Tue, 02 Sep 2008 00:00:00 +0000 /entries/VolpanoSmith.html Arrow and Gibbard-Satterthwaite /entries/ArrowImpossibilityGS.html Mon, 01 Sep 2008 00:00:00 +0000 /entries/ArrowImpossibilityGS.html Fun With Functions /entries/FunWithFunctions.html Tue, 26 Aug 2008 00:00:00 +0000 /entries/FunWithFunctions.html Formal Verification of Modern SAT Solvers /entries/SATSolverVerification.html Wed, 23 Jul 2008 00:00:00 +0000 /entries/SATSolverVerification.html Recursion Theory I /entries/Recursion-Theory-I.html Sat, 05 Apr 2008 00:00:00 +0000 /entries/Recursion-Theory-I.html A Sequential Imperative Programming Language Syntax, Semantics, Hoare Logics and Verification Environment /entries/Simpl.html Fri, 29 Feb 2008 00:00:00 +0000 /entries/Simpl.html BDD Normalisation /entries/BDD.html Fri, 29 Feb 2008 00:00:00 +0000 /entries/BDD.html Normalization by Evaluation /entries/NormByEval.html Mon, 18 Feb 2008 00:00:00 +0000 /entries/NormByEval.html Quantifier Elimination for Linear Arithmetic /entries/LinearQuantifierElim.html Fri, 11 Jan 2008 00:00:00 +0000 /entries/LinearQuantifierElim.html Formalization of Conflict Analysis of Programs with Procedures, Thread Creation, and Monitors /entries/Program-Conflict-Analysis.html Fri, 14 Dec 2007 00:00:00 +0000 /entries/Program-Conflict-Analysis.html Jinja with Threads /entries/JinjaThreads.html Mon, 03 Dec 2007 00:00:00 +0000 /entries/JinjaThreads.html Much Ado About Two /entries/MuchAdoAboutTwo.html Tue, 06 Nov 2007 00:00:00 +0000 /entries/MuchAdoAboutTwo.html Fermat's Last Theorem for Exponents 3 and 4 and the Parametrisation of Pythagorean Triples /entries/Fermat3_4.html Sun, 12 Aug 2007 00:00:00 +0000 /entries/Fermat3_4.html Sums of Two and Four Squares /entries/SumSquares.html Sun, 12 Aug 2007 00:00:00 +0000 /entries/SumSquares.html Fundamental Properties of Valuation Theory and Hensel's Lemma /entries/Valuation.html Wed, 08 Aug 2007 00:00:00 +0000 /entries/Valuation.html First-Order Logic According to Fitting /entries/FOL-Fitting.html Thu, 02 Aug 2007 00:00:00 +0000 /entries/FOL-Fitting.html POPLmark Challenge Via de Bruijn Indices /entries/POPLmark-deBruijn.html Thu, 02 Aug 2007 00:00:00 +0000 /entries/POPLmark-deBruijn.html Hotel Key Card System /entries/HotelKeyCards.html Sat, 09 Sep 2006 00:00:00 +0000 /entries/HotelKeyCards.html Abstract Hoare Logics /entries/Abstract-Hoare-Logics.html Tue, 08 Aug 2006 00:00:00 +0000 /entries/Abstract-Hoare-Logics.html Flyspeck I: Tame Graphs /entries/Flyspeck-Tame.html Mon, 22 May 2006 00:00:00 +0000 /entries/Flyspeck-Tame.html CoreC++ /entries/CoreC++.html Mon, 15 May 2006 00:00:00 +0000 /entries/CoreC++.html A Theory of Featherweight Java in Isabelle/HOL /entries/FeatherweightJava.html Fri, 31 Mar 2006 00:00:00 +0000 /entries/FeatherweightJava.html Instances of Schneider's generalized protocol of clock synchronization /entries/ClockSynchInst.html Wed, 15 Mar 2006 00:00:00 +0000 /entries/ClockSynchInst.html Cauchy's Mean Theorem and the Cauchy-Schwarz Inequality /entries/Cauchy.html Tue, 14 Mar 2006 00:00:00 +0000 /entries/Cauchy.html Countable Ordinals /entries/Ordinal.html Fri, 11 Nov 2005 00:00:00 +0000 /entries/Ordinal.html Fast Fourier Transform /entries/FFT.html Wed, 12 Oct 2005 00:00:00 +0000 /entries/FFT.html Formalization of a Generalized Protocol for Clock Synchronization /entries/GenClock.html Fri, 24 Jun 2005 00:00:00 +0000 /entries/GenClock.html Proving the Correctness of Disk Paxos /entries/DiskPaxos.html Wed, 22 Jun 2005 00:00:00 +0000 /entries/DiskPaxos.html Jive Data and Store Model /entries/JiveDataStoreModel.html Mon, 20 Jun 2005 00:00:00 +0000 /entries/JiveDataStoreModel.html Jinja is not Java /entries/Jinja.html Wed, 01 Jun 2005 00:00:00 +0000 /entries/Jinja.html SHA1, RSA, PSS and more /entries/RSAPSS.html Mon, 02 May 2005 00:00:00 +0000 /entries/RSAPSS.html Category Theory to Yoneda's Lemma /entries/Category.html Thu, 21 Apr 2005 00:00:00 +0000 /entries/Category.html File Refinement /entries/FileRefinement.html Thu, 09 Dec 2004 00:00:00 +0000 /entries/FileRefinement.html Integration theory and random variables /entries/Integration.html Fri, 19 Nov 2004 00:00:00 +0000 /entries/Integration.html A Mechanically Verified, Efficient, Sound and Complete Theorem Prover For First Order Logic /entries/Verified-Prover.html Tue, 28 Sep 2004 00:00:00 +0000 /entries/Verified-Prover.html Completeness theorem /entries/Completeness.html Mon, 20 Sep 2004 00:00:00 +0000 /entries/Completeness.html Ramsey's theorem, infinitary version /entries/Ramsey-Infinite.html Mon, 20 Sep 2004 00:00:00 +0000 /entries/Ramsey-Infinite.html Compiling Exceptions Correctly /entries/Compiling-Exceptions-Correctly.html Fri, 09 Jul 2004 00:00:00 +0000 /entries/Compiling-Exceptions-Correctly.html Depth First Search /entries/Depth-First-Search.html Thu, 24 Jun 2004 00:00:00 +0000 /entries/Depth-First-Search.html Groups, Rings and Modules /entries/Group-Ring-Module.html Tue, 18 May 2004 00:00:00 +0000 /entries/Group-Ring-Module.html Lazy Lists II /entries/Lazy-Lists-II.html Mon, 26 Apr 2004 00:00:00 +0000 /entries/Lazy-Lists-II.html Topology /entries/Topology.html Mon, 26 Apr 2004 00:00:00 +0000 /entries/Topology.html Binary Search Trees /entries/BinarySearchTree.html Mon, 05 Apr 2004 00:00:00 +0000 /entries/BinarySearchTree.html Functional Automata /entries/Functional-Automata.html Tue, 30 Mar 2004 00:00:00 +0000 /entries/Functional-Automata.html AVL Trees /entries/AVL-Trees.html Fri, 19 Mar 2004 00:00:00 +0000 /entries/AVL-Trees.html Mini ML /entries/MiniML.html Fri, 19 Mar 2004 00:00:00 +0000 /entries/MiniML.html Abortable_Linearizable_Modules /theories/abortable_linearizable_modules/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/abortable_linearizable_modules/ About /about/ Mon, 01 Jan 0001 00:00:00 +0000 /about/ The Archive of Formal Proofs is a collection of proof libraries, examples, and larger scientific developments, mechanically checked in the theorem prover Isabelle. It is organized in the way of a scientific journal. Submissions are refereed. The archive repository is hosted on Heptapod to provide easy free access to archive entries. The entries are tested and maintained continuously against the current stable release of Isabelle. Older versions of archive entries will remain available. Abs_Int_ITP2012 /theories/abs_int_itp2012/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/abs_int_itp2012/ Abstract-Hoare-Logics /theories/abstract-hoare-logics/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/abstract-hoare-logics/ Abstract-Rewriting /theories/abstract-rewriting/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/abstract-rewriting/ Abstract_Completeness /theories/abstract_completeness/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/abstract_completeness/ Abstract_Soundness /theories/abstract_soundness/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/abstract_soundness/ Ackermanns_not_PR /theories/ackermanns_not_pr/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ackermanns_not_pr/ Actuarial_Mathematics /theories/actuarial_mathematics/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/actuarial_mathematics/ Adaptive_State_Counting /theories/adaptive_state_counting/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/adaptive_state_counting/ ADS_Functor /theories/ads_functor/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ads_functor/ Affine_Arithmetic /theories/affine_arithmetic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/affine_arithmetic/ Aggregation_Algebras /theories/aggregation_algebras/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/aggregation_algebras/ AI_Planning_Languages_Semantics /theories/ai_planning_languages_semantics/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ai_planning_languages_semantics/ Akra_Bazzi /theories/akra_bazzi/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/akra_bazzi/ Algebraic_Numbers /theories/algebraic_numbers/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/algebraic_numbers/ Algebraic_VCs /theories/algebraic_vcs/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/algebraic_vcs/ Allen_Calculus /theories/allen_calculus/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/allen_calculus/ Amicable_Numbers /theories/amicable_numbers/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/amicable_numbers/ Amortized_Complexity /theories/amortized_complexity/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/amortized_complexity/ AnselmGod /theories/anselmgod/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/anselmgod/ AODV /theories/aodv/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/aodv/ Applicative_Lifting /theories/applicative_lifting/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/applicative_lifting/ Approximation_Algorithms /theories/approximation_algorithms/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/approximation_algorithms/ Architectural_Design_Patterns /theories/architectural_design_patterns/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/architectural_design_patterns/ Aristotles_Assertoric_Syllogistic /theories/aristotles_assertoric_syllogistic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/aristotles_assertoric_syllogistic/ Arith_Prog_Rel_Primes /theories/arith_prog_rel_primes/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/arith_prog_rel_primes/ ArrowImpossibilityGS /theories/arrowimpossibilitygs/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/arrowimpossibilitygs/ Attack_Trees /theories/attack_trees/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/attack_trees/ Auto2_HOL /theories/auto2_hol/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/auto2_hol/ Auto2_Imperative_HOL /theories/auto2_imperative_hol/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/auto2_imperative_hol/ AutoFocus-Stream /theories/autofocus-stream/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/autofocus-stream/ Automated_Stateful_Protocol_Verification /theories/automated_stateful_protocol_verification/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/automated_stateful_protocol_verification/ Automatic_Refinement /theories/automatic_refinement/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/automatic_refinement/ AVL-Trees /theories/avl-trees/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/avl-trees/ AWN /theories/awn/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/awn/ AxiomaticCategoryTheory /theories/axiomaticcategorytheory/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/axiomaticcategorytheory/ Balog_Szemeredi_Gowers /theories/balog_szemeredi_gowers/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/balog_szemeredi_gowers/ Banach_Steinhaus /theories/banach_steinhaus/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/banach_steinhaus/ BD_Security_Compositional /theories/bd_security_compositional/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/bd_security_compositional/ BDD /theories/bdd/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/bdd/ Belief_Revision /theories/belief_revision/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/belief_revision/ Bell_Numbers_Spivey /theories/bell_numbers_spivey/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/bell_numbers_spivey/ BenOr_Kozen_Reif /theories/benor_kozen_reif/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/benor_kozen_reif/ Berlekamp_Zassenhaus /theories/berlekamp_zassenhaus/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/berlekamp_zassenhaus/ Bernoulli /theories/bernoulli/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/bernoulli/ Bertrands_Postulate /theories/bertrands_postulate/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/bertrands_postulate/ Bicategory /theories/bicategory/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/bicategory/ BinarySearchTree /theories/binarysearchtree/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/binarysearchtree/ Binding_Syntax_Theory /theories/binding_syntax_theory/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/binding_syntax_theory/ Binomial-Heaps /theories/binomial-heaps/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/binomial-heaps/ Binomial-Queues /theories/binomial-queues/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/binomial-queues/ BirdKMP /theories/birdkmp/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/birdkmp/ Blue_Eyes /theories/blue_eyes/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/blue_eyes/ BNF_CC /theories/bnf_cc/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/bnf_cc/ BNF_Operations /theories/bnf_operations/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/bnf_operations/ Bondy /theories/bondy/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/bondy/ Boolean_Expression_Checkers /theories/boolean_expression_checkers/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/boolean_expression_checkers/ Boolos_Curious_Inference /theories/boolos_curious_inference/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/boolos_curious_inference/ Bounded_Deducibility_Security /theories/bounded_deducibility_security/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/bounded_deducibility_security/ BTree /theories/btree/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/btree/ Buchi_Complementation /theories/buchi_complementation/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/buchi_complementation/ Budan_Fourier /theories/budan_fourier/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/budan_fourier/ Buffons_Needle /theories/buffons_needle/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/buffons_needle/ Buildings /theories/buildings/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/buildings/ BytecodeLogicJmlTypes /theories/bytecodelogicjmltypes/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/bytecodelogicjmltypes/ C2KA_DistributedSystems /theories/c2ka_distributedsystems/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/c2ka_distributedsystems/ CakeML /theories/cakeml/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cakeml/ CakeML_Codegen /theories/cakeml_codegen/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cakeml_codegen/ Call_Arity /theories/call_arity/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/call_arity/ Card_Equiv_Relations /theories/card_equiv_relations/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/card_equiv_relations/ Card_Multisets /theories/card_multisets/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/card_multisets/ Card_Number_Partitions /theories/card_number_partitions/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/card_number_partitions/ Card_Partitions /theories/card_partitions/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/card_partitions/ Cartan_FP /theories/cartan_fp/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cartan_fp/ Case_Labeling /theories/case_labeling/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/case_labeling/ Catalan_Numbers /theories/catalan_numbers/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/catalan_numbers/ Category /theories/category/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/category/ Category2 /theories/category2/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/category2/ Category3 /theories/category3/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/category3/ Cauchy /theories/cauchy/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cauchy/ CAVA_Automata /theories/cava_automata/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cava_automata/ CAVA_Base /theories/cava_base/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cava_base/ CAVA_LTL_Modelchecker /theories/cava_ltl_modelchecker/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cava_ltl_modelchecker/ CAVA_Setup /theories/cava_setup/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cava_setup/ Cayley_Hamilton /theories/cayley_hamilton/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cayley_hamilton/ CCS /theories/ccs/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ccs/ Certification_Monads /theories/certification_monads/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/certification_monads/ Chandy_Lamport /theories/chandy_lamport/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/chandy_lamport/ Chord_Segments /theories/chord_segments/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/chord_segments/ Circus /theories/circus/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/circus/ CISC-Kernel /theories/cisc-kernel/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cisc-kernel/ Clean /theories/clean/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/clean/ Clique_and_Monotone_Circuits /theories/clique_and_monotone_circuits/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/clique_and_monotone_circuits/ ClockSynchInst /theories/clocksynchinst/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/clocksynchinst/ Closest_Pair_Points /theories/closest_pair_points/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/closest_pair_points/ CoCon /theories/cocon/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cocon/ CofGroups /theories/cofgroups/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cofgroups/ Coinductive /theories/coinductive/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/coinductive/ Coinductive_Languages /theories/coinductive_languages/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/coinductive_languages/ Collections /theories/collections/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/collections/ Collections_Examples /theories/collections_examples/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/collections_examples/ Combinable_Wands /theories/combinable_wands/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/combinable_wands/ Combinatorial_Enumeration_Algorithms /theories/combinatorial_enumeration_algorithms/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/combinatorial_enumeration_algorithms/ Combinatorics_Words /theories/combinatorics_words/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/combinatorics_words/ Combinatorics_Words_Graph_Lemma /theories/combinatorics_words_graph_lemma/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/combinatorics_words_graph_lemma/ Combinatorics_Words_Lyndon /theories/combinatorics_words_lyndon/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/combinatorics_words_lyndon/ Commuting_Hermitian /theories/commuting_hermitian/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/commuting_hermitian/ Comparison_Sort_Lower_Bound /theories/comparison_sort_lower_bound/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/comparison_sort_lower_bound/ Compiling-Exceptions-Correctly /theories/compiling-exceptions-correctly/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/compiling-exceptions-correctly/ Complete_Non_Orders /theories/complete_non_orders/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/complete_non_orders/ Completeness /theories/completeness/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/completeness/ Complex_Bounded_Operators /theories/complex_bounded_operators/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/complex_bounded_operators/ Complex_Geometry /theories/complex_geometry/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/complex_geometry/ Complx /theories/complx/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/complx/ ComponentDependencies /theories/componentdependencies/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/componentdependencies/ Concurrent_Ref_Alg /theories/concurrent_ref_alg/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/concurrent_ref_alg/ Concurrent_Revisions /theories/concurrent_revisions/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/concurrent_revisions/ ConcurrentGC /theories/concurrentgc/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/concurrentgc/ ConcurrentIMP /theories/concurrentimp/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/concurrentimp/ Conditional_Simplification /theories/conditional_simplification/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/conditional_simplification/ Conditional_Transfer_Rule /theories/conditional_transfer_rule/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/conditional_transfer_rule/ Consensus_Refined /theories/consensus_refined/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/consensus_refined/ Constructive_Cryptography /theories/constructive_cryptography/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/constructive_cryptography/ Constructive_Cryptography_CM /theories/constructive_cryptography_cm/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/constructive_cryptography_cm/ Constructor_Funs /theories/constructor_funs/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/constructor_funs/ Containers /theories/containers/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/containers/ Containers-Benchmarks /theories/containers-benchmarks/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/containers-benchmarks/ Core_DOM /theories/core_dom/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/core_dom/ Core_SC_DOM /theories/core_sc_dom/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/core_sc_dom/ CoreC++ /theories/corec++/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/corec++/ Correctness_Algebras /theories/correctness_algebras/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/correctness_algebras/ CoSMed /theories/cosmed/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cosmed/ CoSMeDis /theories/cosmedis/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cosmedis/ Cotangent_PFD_Formula /theories/cotangent_pfd_formula/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cotangent_pfd_formula/ Count_Complex_Roots /theories/count_complex_roots/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/count_complex_roots/ CRDT /theories/crdt/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/crdt/ CryptHOL /theories/crypthol/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/crypthol/ CryptoBasedCompositionalProperties /theories/cryptobasedcompositionalproperties/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cryptobasedcompositionalproperties/ CRYSTALS-Kyber /theories/crystals-kyber/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/crystals-kyber/ CSP_RefTK /theories/csp_reftk/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/csp_reftk/ Cubic_Quartic_Equations /theories/cubic_quartic_equations/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cubic_quartic_equations/ CYK /theories/cyk/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/cyk/ CZH_Elementary_Categories /theories/czh_elementary_categories/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/czh_elementary_categories/ CZH_Foundations /theories/czh_foundations/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/czh_foundations/ CZH_Universal_Constructions /theories/czh_universal_constructions/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/czh_universal_constructions/ DataRefinementIBP /theories/datarefinementibp/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/datarefinementibp/ Datatype_Order_Generator /theories/datatype_order_generator/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/datatype_order_generator/ Decl_Sem_Fun_PL /theories/decl_sem_fun_pl/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/decl_sem_fun_pl/ Decreasing-Diagrams /theories/decreasing-diagrams/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/decreasing-diagrams/ Decreasing-Diagrams-II /theories/decreasing-diagrams-ii/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/decreasing-diagrams-ii/ Dedekind_Real /theories/dedekind_real/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/dedekind_real/ Deep_Learning /theories/deep_learning/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/deep_learning/ Delta_System_Lemma /theories/delta_system_lemma/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/delta_system_lemma/ Density_Compiler /theories/density_compiler/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/density_compiler/ Dependent_SIFUM_Refinement /theories/dependent_sifum_refinement/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/dependent_sifum_refinement/ Dependent_SIFUM_Type_Systems /theories/dependent_sifum_type_systems/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/dependent_sifum_type_systems/ Depth-First-Search /theories/depth-first-search/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/depth-first-search/ Derangements /theories/derangements/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/derangements/ Deriving /theories/deriving/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/deriving/ Descartes_Sign_Rule /theories/descartes_sign_rule/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/descartes_sign_rule/ Design_Theory /theories/design_theory/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/design_theory/ DFS_Framework /theories/dfs_framework/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/dfs_framework/ Dict_Construction /theories/dict_construction/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/dict_construction/ Differential_Dynamic_Logic /theories/differential_dynamic_logic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/differential_dynamic_logic/ Differential_Game_Logic /theories/differential_game_logic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/differential_game_logic/ Digit_Expansions /theories/digit_expansions/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/digit_expansions/ Dijkstra_Shortest_Path /theories/dijkstra_shortest_path/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/dijkstra_shortest_path/ Diophantine_Eqns_Lin_Hom /theories/diophantine_eqns_lin_hom/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/diophantine_eqns_lin_hom/ Dirichlet_L /theories/dirichlet_l/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/dirichlet_l/ Dirichlet_Series /theories/dirichlet_series/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/dirichlet_series/ Discrete_Summation /theories/discrete_summation/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/discrete_summation/ DiscretePricing /theories/discretepricing/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/discretepricing/ DiskPaxos /theories/diskpaxos/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/diskpaxos/ DOM_Components /theories/dom_components/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/dom_components/ Dominance_CHK /theories/dominance_chk/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/dominance_chk/ Download the Archive /download/ Mon, 01 Jan 0001 00:00:00 +0000 /download/ Current stable version (for most recent Isabelle release): Download all sessions: afp-current.tar.gz (~70 MB) Older stable versions: Please use the sourceforge download system to access older versions of the archive. Mercurial access: The AFP repositories with tooling and metadata can be found at Heptapod. In particular, the development version of the Archive (for the development version of Isabelle) is available there. How to refer to AFP entries: You can refer to AFP entries by using the AFP as an Isabelle component. DPRM_Theorem /theories/dprm_theorem/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/dprm_theorem/ DPT-SAT-Solver /theories/dpt-sat-solver/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/dpt-sat-solver/ Dynamic_Tables /theories/dynamic_tables/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/dynamic_tables/ DynamicArchitectures /theories/dynamicarchitectures/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/dynamicarchitectures/ E_Transcendental /theories/e_transcendental/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/e_transcendental/ Echelon_Form /theories/echelon_form/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/echelon_form/ EdmondsKarp_Maxflow /theories/edmondskarp_maxflow/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/edmondskarp_maxflow/ Efficient-Mergesort /theories/efficient-mergesort/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/efficient-mergesort/ Elliptic_Curves_Group_Law /theories/elliptic_curves_group_law/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/elliptic_curves_group_law/ Encodability_Process_Calculi /theories/encodability_process_calculi/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/encodability_process_calculi/ Entry Submission /submission/ Mon, 01 Jan 0001 00:00:00 +0000 /submission/ Submission Guidelines The submission must follow the following Isabelle style rules. For additional guidelines on Isabelle proofs, also see the this guide (feel free to follow all of these; only the below are mandatory). Technical details about the submission process and the format of the submission are explained on the submission site. No use of the commands sorry or back. Instantiations must not use Isabelle-generated names such as xa — use Isar, the subgoal command or rename_tac to avoid such names. Epistemic_Logic /theories/epistemic_logic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/epistemic_logic/ Equivalence_Relation_Enumeration /theories/equivalence_relation_enumeration/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/equivalence_relation_enumeration/ Ergodic_Theory /theories/ergodic_theory/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ergodic_theory/ Error_Function /theories/error_function/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/error_function/ Euler_MacLaurin /theories/euler_maclaurin/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/euler_maclaurin/ Euler_Partition /theories/euler_partition/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/euler_partition/ Eval_FO /theories/eval_fo/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/eval_fo/ Extended_Finite_State_Machine_Inference /theories/extended_finite_state_machine_inference/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/extended_finite_state_machine_inference/ Extended_Finite_State_Machines /theories/extended_finite_state_machines/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/extended_finite_state_machines/ Factor_Algebraic_Polynomial /theories/factor_algebraic_polynomial/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/factor_algebraic_polynomial/ Factored_Transition_System_Bounding /theories/factored_transition_system_bounding/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/factored_transition_system_bounding/ Falling_Factorial_Sum /theories/falling_factorial_sum/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/falling_factorial_sum/ Farkas /theories/farkas/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/farkas/ Featherweight_OCL /theories/featherweight_ocl/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/featherweight_ocl/ FeatherweightJava /theories/featherweightjava/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/featherweightjava/ Fermat3_4 /theories/fermat3_4/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/fermat3_4/ FFT /theories/fft/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/fft/ FileRefinement /theories/filerefinement/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/filerefinement/ FinFun /theories/finfun/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/finfun/ Finger-Trees /theories/finger-trees/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/finger-trees/ Finite-Map-Extras /theories/finite-map-extras/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/finite-map-extras/ Finite_Automata_HF /theories/finite_automata_hf/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/finite_automata_hf/ Finite_Fields /theories/finite_fields/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/finite_fields/ Finitely_Generated_Abelian_Groups /theories/finitely_generated_abelian_groups/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/finitely_generated_abelian_groups/ First_Order_Terms /theories/first_order_terms/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/first_order_terms/ First_Welfare_Theorem /theories/first_welfare_theorem/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/first_welfare_theorem/ Fishburn_Impossibility /theories/fishburn_impossibility/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/fishburn_impossibility/ Fisher_Yates /theories/fisher_yates/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/fisher_yates/ Fishers_Inequality /theories/fishers_inequality/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/fishers_inequality/ Flow_Networks /theories/flow_networks/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/flow_networks/ Floyd_Warshall /theories/floyd_warshall/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/floyd_warshall/ FLP /theories/flp/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/flp/ Flyspeck-Tame /theories/flyspeck-tame/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/flyspeck-tame/ Flyspeck-Tame-Computation /theories/flyspeck-tame-computation/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/flyspeck-tame-computation/ FO_Theory_Rewriting /theories/fo_theory_rewriting/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/fo_theory_rewriting/ FocusStreamsCaseStudies /theories/focusstreamscasestudies/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/focusstreamscasestudies/ FOL-Fitting /theories/fol-fitting/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/fol-fitting/ FOL_Axiomatic /theories/fol_axiomatic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/fol_axiomatic/ FOL_Harrison /theories/fol_harrison/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/fol_harrison/ FOL_Seq_Calc1 /theories/fol_seq_calc1/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/fol_seq_calc1/ FOL_Seq_Calc2 /theories/fol_seq_calc2/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/fol_seq_calc2/ FOL_Seq_Calc3 /theories/fol_seq_calc3/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/fol_seq_calc3/ Forcing /theories/forcing/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/forcing/ Formal_Puiseux_Series /theories/formal_puiseux_series/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/formal_puiseux_series/ Formal_SSA /theories/formal_ssa/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/formal_ssa/ Formula_Derivatives /theories/formula_derivatives/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/formula_derivatives/ Formula_Derivatives-Examples /theories/formula_derivatives-examples/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/formula_derivatives-examples/ Foundation_of_geometry /theories/foundation_of_geometry/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/foundation_of_geometry/ Fourier /theories/fourier/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/fourier/ Free-Boolean-Algebra /theories/free-boolean-algebra/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/free-boolean-algebra/ Free-Groups /theories/free-groups/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/free-groups/ Frequency_Moments /theories/frequency_moments/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/frequency_moments/ Fresh_Identifiers /theories/fresh_identifiers/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/fresh_identifiers/ FSM_Tests /theories/fsm_tests/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/fsm_tests/ Functional-Automata /theories/functional-automata/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/functional-automata/ Functional_Ordered_Resolution_Prover /theories/functional_ordered_resolution_prover/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/functional_ordered_resolution_prover/ FunWithFunctions /theories/funwithfunctions/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/funwithfunctions/ FunWithTilings /theories/funwithtilings/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/funwithtilings/ Furstenberg_Topology /theories/furstenberg_topology/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/furstenberg_topology/ Gabow_SCC /theories/gabow_scc/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/gabow_scc/ Gale_Shapley /theories/gale_shapley/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/gale_shapley/ GaleStewart_Games /theories/galestewart_games/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/galestewart_games/ Game_Based_Crypto /theories/game_based_crypto/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/game_based_crypto/ Gauss-Jordan-Elim-Fun /theories/gauss-jordan-elim-fun/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/gauss-jordan-elim-fun/ Gauss_Jordan /theories/gauss_jordan/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/gauss_jordan/ Gauss_Sums /theories/gauss_sums/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/gauss_sums/ Gaussian_Integers /theories/gaussian_integers/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/gaussian_integers/ GenClock /theories/genclock/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/genclock/ General-Triangle /theories/general-triangle/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/general-triangle/ Generalized_Counting_Sort /theories/generalized_counting_sort/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/generalized_counting_sort/ Generic_Deriving /theories/generic_deriving/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/generic_deriving/ Generic_Join /theories/generic_join/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/generic_join/ GewirthPGCProof /theories/gewirthpgcproof/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/gewirthpgcproof/ Girth_Chromatic /theories/girth_chromatic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/girth_chromatic/ Goedel_HFSet_Semantic /theories/goedel_hfset_semantic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/goedel_hfset_semantic/ Goedel_HFSet_Semanticless /theories/goedel_hfset_semanticless/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/goedel_hfset_semanticless/ Goedel_Incompleteness /theories/goedel_incompleteness/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/goedel_incompleteness/ GoedelGod /theories/goedelgod/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/goedelgod/ Goodstein_Lambda /theories/goodstein_lambda/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/goodstein_lambda/ GPU_Kernel_PL /theories/gpu_kernel_pl/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/gpu_kernel_pl/ Graph_Saturation /theories/graph_saturation/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/graph_saturation/ Graph_Theory /theories/graph_theory/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/graph_theory/ GraphMarkingIBP /theories/graphmarkingibp/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/graphmarkingibp/ Green /theories/green/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/green/ Groebner_Bases /theories/groebner_bases/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/groebner_bases/ Groebner_Macaulay /theories/groebner_macaulay/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/groebner_macaulay/ Gromov_Hyperbolicity /theories/gromov_hyperbolicity/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/gromov_hyperbolicity/ Grothendieck_Schemes /theories/grothendieck_schemes/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/grothendieck_schemes/ Group-Ring-Module /theories/group-ring-module/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/group-ring-module/ Hahn_Jordan_Decomposition /theories/hahn_jordan_decomposition/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hahn_jordan_decomposition/ Hales_Jewett /theories/hales_jewett/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hales_jewett/ Heard_Of /theories/heard_of/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/heard_of/ Hello_World /theories/hello_world/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hello_world/ Help /help/ Mon, 01 Jan 0001 00:00:00 +0000 /help/ This section focuses on the Archive of Formal Proofs. For help with Isabelle, see the Isabelle Documentation. More resources are listed in the Isabelle Quick Access Links. Referring to AFP Entries in Isabelle/JEdit Once you have downloaded the AFP, you can include its articles and theories in your own developments. If you would like to make your work available to others without having to include the AFP articles you depend on, here is how to do it. HereditarilyFinite /theories/hereditarilyfinite/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hereditarilyfinite/ Hermite /theories/hermite/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hermite/ Hermite_Lindemann /theories/hermite_lindemann/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hermite_lindemann/ Hidden_Markov_Models /theories/hidden_markov_models/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hidden_markov_models/ Higher_Order_Terms /theories/higher_order_terms/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/higher_order_terms/ Hoare_Time /theories/hoare_time/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hoare_time/ HOL-CSP /theories/hol-csp/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hol-csp/ HOL-ODE-ARCH-COMP /theories/hol-ode-arch-comp/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hol-ode-arch-comp/ HOL-ODE-Examples /theories/hol-ode-examples/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hol-ode-examples/ HOL-ODE-Numerics /theories/hol-ode-numerics/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hol-ode-numerics/ HOLCF-Prelude /theories/holcf-prelude/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/holcf-prelude/ Hood_Melville_Queue /theories/hood_melville_queue/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hood_melville_queue/ HotelKeyCards /theories/hotelkeycards/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hotelkeycards/ HRB-Slicing /theories/hrb-slicing/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hrb-slicing/ Huffman /theories/huffman/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/huffman/ Hybrid_Logic /theories/hybrid_logic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hybrid_logic/ Hybrid_Multi_Lane_Spatial_Logic /theories/hybrid_multi_lane_spatial_logic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hybrid_multi_lane_spatial_logic/ Hybrid_Systems_VCs /theories/hybrid_systems_vcs/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hybrid_systems_vcs/ HyperCTL /theories/hyperctl/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hyperctl/ Hyperdual /theories/hyperdual/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/hyperdual/ IEEE_Floating_Point /theories/ieee_floating_point/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ieee_floating_point/ IFC_Tracking /theories/ifc_tracking/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ifc_tracking/ IMAP-CRDT /theories/imap-crdt/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/imap-crdt/ IMO2019 /theories/imo2019/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/imo2019/ IMP2 /theories/imp2/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/imp2/ IMP2_Binary_Heap /theories/imp2_binary_heap/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/imp2_binary_heap/ IMP_Compiler /theories/imp_compiler/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/imp_compiler/ IMP_Compiler_Reuse /theories/imp_compiler_reuse/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/imp_compiler_reuse/ Imperative_Insertion_Sort /theories/imperative_insertion_sort/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/imperative_insertion_sort/ Implicational_Logic /theories/implicational_logic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/implicational_logic/ Impossible_Geometry /theories/impossible_geometry/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/impossible_geometry/ Incompleteness /theories/incompleteness/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/incompleteness/ Incredible_Proof_Machine /theories/incredible_proof_machine/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/incredible_proof_machine/ Independence_CH /theories/independence_ch/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/independence_ch/ Inductive_Confidentiality /theories/inductive_confidentiality/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/inductive_confidentiality/ Inductive_Inference /theories/inductive_inference/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/inductive_inference/ InformationFlowSlicing /theories/informationflowslicing/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/informationflowslicing/ InformationFlowSlicing_Inter /theories/informationflowslicing_inter/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/informationflowslicing_inter/ InfPathElimination /theories/infpathelimination/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/infpathelimination/ Integration /theories/integration/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/integration/ Interpolation_Polynomials_HOL_Algebra /theories/interpolation_polynomials_hol_algebra/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/interpolation_polynomials_hol_algebra/ Interpreter_Optimizations /theories/interpreter_optimizations/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/interpreter_optimizations/ Interval_Arithmetic_Word32 /theories/interval_arithmetic_word32/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/interval_arithmetic_word32/ Intro_Dest_Elim /theories/intro_dest_elim/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/intro_dest_elim/ Involutions2Squares /theories/involutions2squares/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/involutions2squares/ IP_Addresses /theories/ip_addresses/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ip_addresses/ Iptables_Semantics /theories/iptables_semantics/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/iptables_semantics/ Iptables_Semantics_Examples /theories/iptables_semantics_examples/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/iptables_semantics_examples/ Iptables_Semantics_Examples_Big /theories/iptables_semantics_examples_big/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/iptables_semantics_examples_big/ Irrational_Series_Erdos_Straus /theories/irrational_series_erdos_straus/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/irrational_series_erdos_straus/ Irrationality_J_Hancl /theories/irrationality_j_hancl/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/irrationality_j_hancl/ Irrationals_From_THEBOOK /theories/irrationals_from_thebook/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/irrationals_from_thebook/ Isabelle_C /theories/isabelle_c/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/isabelle_c/ Isabelle_Marries_Dirac /theories/isabelle_marries_dirac/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/isabelle_marries_dirac/ Isabelle_Meta_Model /theories/isabelle_meta_model/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/isabelle_meta_model/ IsaGeoCoq /theories/isageocoq/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/isageocoq/ IsaNet /theories/isanet/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/isanet/ Jacobson_Basic_Algebra /theories/jacobson_basic_algebra/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/jacobson_basic_algebra/ Jinja /theories/jinja/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/jinja/ JinjaDCI /theories/jinjadci/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/jinjadci/ JinjaThreads /theories/jinjathreads/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/jinjathreads/ JiveDataStoreModel /theories/jivedatastoremodel/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/jivedatastoremodel/ Jordan_Hoelder /theories/jordan_hoelder/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/jordan_hoelder/ Jordan_Normal_Form /theories/jordan_normal_form/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/jordan_normal_form/ KAD /theories/kad/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/kad/ KAT_and_DRA /theories/kat_and_dra/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/kat_and_dra/ KBPs /theories/kbps/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/kbps/ KD_Tree /theories/kd_tree/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/kd_tree/ Key_Agreement_Strong_Adversaries /theories/key_agreement_strong_adversaries/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/key_agreement_strong_adversaries/ Khovanskii_Theorem /theories/khovanskii_theorem/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/khovanskii_theorem/ Kleene_Algebra /theories/kleene_algebra/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/kleene_algebra/ Knights_Tour /theories/knights_tour/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/knights_tour/ Knot_Theory /theories/knot_theory/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/knot_theory/ Knuth_Bendix_Order /theories/knuth_bendix_order/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/knuth_bendix_order/ Knuth_Morris_Pratt /theories/knuth_morris_pratt/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/knuth_morris_pratt/ Koenigsberg_Friendship /theories/koenigsberg_friendship/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/koenigsberg_friendship/ Kruskal /theories/kruskal/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/kruskal/ Kuratowski_Closure_Complement /theories/kuratowski_closure_complement/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/kuratowski_closure_complement/ Lam-ml-Normalization /theories/lam-ml-normalization/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lam-ml-normalization/ Lambda_Free_EPO /theories/lambda_free_epo/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lambda_free_epo/ Lambda_Free_KBOs /theories/lambda_free_kbos/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lambda_free_kbos/ Lambda_Free_RPOs /theories/lambda_free_rpos/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lambda_free_rpos/ LambdaAuth /theories/lambdaauth/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lambdaauth/ LambdaMu /theories/lambdamu/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lambdamu/ Lambert_W /theories/lambert_w/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lambert_w/ Landau_Symbols /theories/landau_symbols/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/landau_symbols/ Laplace_Transform /theories/laplace_transform/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/laplace_transform/ Latin_Square /theories/latin_square/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/latin_square/ LatticeProperties /theories/latticeproperties/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/latticeproperties/ Launchbury /theories/launchbury/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/launchbury/ Laws_of_Large_Numbers /theories/laws_of_large_numbers/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/laws_of_large_numbers/ Lazy-Lists-II /theories/lazy-lists-ii/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lazy-lists-ii/ Lazy_Case /theories/lazy_case/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lazy_case/ Lehmer /theories/lehmer/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lehmer/ LEM /theories/lem/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lem/ Lifting_Definition_Option /theories/lifting_definition_option/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lifting_definition_option/ Lifting_the_Exponent /theories/lifting_the_exponent/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lifting_the_exponent/ LightweightJava /theories/lightweightjava/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lightweightjava/ Linear_Inequalities /theories/linear_inequalities/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/linear_inequalities/ Linear_Programming /theories/linear_programming/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/linear_programming/ Linear_Recurrences /theories/linear_recurrences/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/linear_recurrences/ Linear_Recurrences_Solver /theories/linear_recurrences_solver/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/linear_recurrences_solver/ LinearQuantifierElim /theories/linearquantifierelim/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/linearquantifierelim/ Liouville_Numbers /theories/liouville_numbers/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/liouville_numbers/ List-Index /theories/list-index/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/list-index/ List-Infinite /theories/list-infinite/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/list-infinite/ List_Interleaving /theories/list_interleaving/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/list_interleaving/ List_Inversions /theories/list_inversions/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/list_inversions/ List_Update /theories/list_update/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/list_update/ LLL_Basis_Reduction /theories/lll_basis_reduction/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lll_basis_reduction/ LLL_Factorization /theories/lll_factorization/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lll_factorization/ Localization_Ring /theories/localization_ring/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/localization_ring/ LocalLexing /theories/locallexing/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/locallexing/ Locally-Nameless-Sigma /theories/locally-nameless-sigma/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/locally-nameless-sigma/ LOFT /theories/loft/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/loft/ Logging_Independent_Anonymity /theories/logging_independent_anonymity/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/logging_independent_anonymity/ Lorenz_Approximation /theories/lorenz_approximation/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lorenz_approximation/ Lorenz_C0 /theories/lorenz_c0/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lorenz_c0/ Lorenz_C1 /theories/lorenz_c1/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lorenz_c1/ Lowe_Ontological_Argument /theories/lowe_ontological_argument/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lowe_ontological_argument/ Lower_Semicontinuous /theories/lower_semicontinuous/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lower_semicontinuous/ Lp /theories/lp/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lp/ LP_Duality /theories/lp_duality/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lp_duality/ LTL /theories/ltl/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ltl/ LTL_Master_Theorem /theories/ltl_master_theorem/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ltl_master_theorem/ LTL_Normal_Form /theories/ltl_normal_form/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ltl_normal_form/ LTL_to_DRA /theories/ltl_to_dra/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ltl_to_dra/ LTL_to_GBA /theories/ltl_to_gba/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ltl_to_gba/ Lucas_Theorem /theories/lucas_theorem/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/lucas_theorem/ Markov_Models /theories/markov_models/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/markov_models/ Marriage /theories/marriage/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/marriage/ Mason_Stothers /theories/mason_stothers/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/mason_stothers/ Matrices_for_ODEs /theories/matrices_for_odes/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/matrices_for_odes/ Matrix /theories/matrix/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/matrix/ Matrix_Tensor /theories/matrix_tensor/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/matrix_tensor/ Matroids /theories/matroids/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/matroids/ Max-Card-Matching /theories/max-card-matching/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/max-card-matching/ Maximum_Segment_Sum /theories/maximum_segment_sum/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/maximum_segment_sum/ MDP-Algorithms /theories/mdp-algorithms/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/mdp-algorithms/ MDP-Rewards /theories/mdp-rewards/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/mdp-rewards/ Median_Method /theories/median_method/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/median_method/ Median_Of_Medians_Selection /theories/median_of_medians_selection/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/median_of_medians_selection/ Menger /theories/menger/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/menger/ Mereology /theories/mereology/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/mereology/ Mersenne_Primes /theories/mersenne_primes/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/mersenne_primes/ Metalogic_ProofChecker /theories/metalogic_proofchecker/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/metalogic_proofchecker/ MFMC_Countable /theories/mfmc_countable/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/mfmc_countable/ MFODL_Monitor_Optimized /theories/mfodl_monitor_optimized/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/mfodl_monitor_optimized/ MFOTL_Monitor /theories/mfotl_monitor/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/mfotl_monitor/ Minimal_SSA /theories/minimal_ssa/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/minimal_ssa/ MiniML /theories/miniml/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/miniml/ MiniSail /theories/minisail/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/minisail/ Minkowskis_Theorem /theories/minkowskis_theorem/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/minkowskis_theorem/ Minsky_Machines /theories/minsky_machines/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/minsky_machines/ Modal_Logics_for_NTS /theories/modal_logics_for_nts/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/modal_logics_for_nts/ Modular_arithmetic_LLL_and_HNF_algorithms /theories/modular_arithmetic_lll_and_hnf_algorithms/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/modular_arithmetic_lll_and_hnf_algorithms/ Modular_Assembly_Kit_Security /theories/modular_assembly_kit_security/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/modular_assembly_kit_security/ Monad_Memo_DP /theories/monad_memo_dp/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/monad_memo_dp/ Monad_Normalisation /theories/monad_normalisation/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/monad_normalisation/ MonoBoolTranAlgebra /theories/monobooltranalgebra/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/monobooltranalgebra/ MonoidalCategory /theories/monoidalcategory/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/monoidalcategory/ Monomorphic_Monad /theories/monomorphic_monad/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/monomorphic_monad/ MSO_Regex_Equivalence /theories/mso_regex_equivalence/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/mso_regex_equivalence/ MuchAdoAboutTwo /theories/muchadoabouttwo/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/muchadoabouttwo/ Multi_Party_Computation /theories/multi_party_computation/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/multi_party_computation/ Multirelations /theories/multirelations/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/multirelations/ Multiset_Ordering_NPC /theories/multiset_ordering_npc/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/multiset_ordering_npc/ Myhill-Nerode /theories/myhill-nerode/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/myhill-nerode/ Name_Carrying_Type_Inference /theories/name_carrying_type_inference/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/name_carrying_type_inference/ Nano_JSON /theories/nano_json/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/nano_json/ Nash_Williams /theories/nash_williams/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/nash_williams/ Nat-Interval-Logic /theories/nat-interval-logic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/nat-interval-logic/ Native_Word /theories/native_word/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/native_word/ Nested_Multisets_Ordinals /theories/nested_multisets_ordinals/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/nested_multisets_ordinals/ Network_Security_Policy_Verification /theories/network_security_policy_verification/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/network_security_policy_verification/ Neumann_Morgenstern_Utility /theories/neumann_morgenstern_utility/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/neumann_morgenstern_utility/ No_FTL_observers /theories/no_ftl_observers/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/no_ftl_observers/ Nominal2 /theories/nominal2/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/nominal2/ Noninterference_Concurrent_Composition /theories/noninterference_concurrent_composition/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/noninterference_concurrent_composition/ Noninterference_CSP /theories/noninterference_csp/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/noninterference_csp/ Noninterference_Generic_Unwinding /theories/noninterference_generic_unwinding/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/noninterference_generic_unwinding/ Noninterference_Inductive_Unwinding /theories/noninterference_inductive_unwinding/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/noninterference_inductive_unwinding/ Noninterference_Ipurge_Unwinding /theories/noninterference_ipurge_unwinding/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/noninterference_ipurge_unwinding/ Noninterference_Sequential_Composition /theories/noninterference_sequential_composition/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/noninterference_sequential_composition/ NormByEval /theories/normbyeval/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/normbyeval/ Nullstellensatz /theories/nullstellensatz/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/nullstellensatz/ Number_Theoretic_Transform /theories/number_theoretic_transform/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/number_theoretic_transform/ Octonions /theories/octonions/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/octonions/ Old_Datatype_Show /theories/old_datatype_show/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/old_datatype_show/ Open_Induction /theories/open_induction/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/open_induction/ OpSets /theories/opsets/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/opsets/ Optics /theories/optics/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/optics/ Optimal_BST /theories/optimal_bst/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/optimal_bst/ Orbit_Stabiliser /theories/orbit_stabiliser/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/orbit_stabiliser/ Order_Lattice_Props /theories/order_lattice_props/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/order_lattice_props/ Ordered_Resolution_Prover /theories/ordered_resolution_prover/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ordered_resolution_prover/ Ordinal /theories/ordinal/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ordinal/ Ordinal_Partitions /theories/ordinal_partitions/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ordinal_partitions/ Ordinals_and_Cardinals /theories/ordinals_and_cardinals/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ordinals_and_cardinals/ Ordinary_Differential_Equations /theories/ordinary_differential_equations/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ordinary_differential_equations/ PAC_Checker /theories/pac_checker/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/pac_checker/ Package_logic /theories/package_logic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/package_logic/ Padic_Field /theories/padic_field/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/padic_field/ Padic_Ints /theories/padic_ints/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/padic_ints/ Pairing_Heap /theories/pairing_heap/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/pairing_heap/ PAL /theories/pal/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/pal/ PAPP_Impossibility /theories/papp_impossibility/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/papp_impossibility/ Paraconsistency /theories/paraconsistency/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/paraconsistency/ Parity_Game /theories/parity_game/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/parity_game/ Partial_Function_MR /theories/partial_function_mr/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/partial_function_mr/ Partial_Order_Reduction /theories/partial_order_reduction/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/partial_order_reduction/ Password_Authentication_Protocol /theories/password_authentication_protocol/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/password_authentication_protocol/ PCF /theories/pcf/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/pcf/ Pell /theories/pell/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/pell/ Perfect-Number-Thm /theories/perfect-number-thm/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/perfect-number-thm/ Perron_Frobenius /theories/perron_frobenius/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/perron_frobenius/ pGCL /theories/pgcl/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/pgcl/ Physical_Quantities /theories/physical_quantities/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/physical_quantities/ Pi_Calculus /theories/pi_calculus/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/pi_calculus/ Pi_Transcendental /theories/pi_transcendental/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/pi_transcendental/ Planarity_Certificates /theories/planarity_certificates/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/planarity_certificates/ PLM /theories/plm/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/plm/ Pluennecke_Ruzsa_Inequality /theories/pluennecke_ruzsa_inequality/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/pluennecke_ruzsa_inequality/ Poincare_Bendixson /theories/poincare_bendixson/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/poincare_bendixson/ Poincare_Disc /theories/poincare_disc/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/poincare_disc/ Polynomial_Factorization /theories/polynomial_factorization/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/polynomial_factorization/ Polynomial_Interpolation /theories/polynomial_interpolation/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/polynomial_interpolation/ Polynomials /theories/polynomials/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/polynomials/ Pop_Refinement /theories/pop_refinement/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/pop_refinement/ POPLmark-deBruijn /theories/poplmark-debruijn/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/poplmark-debruijn/ Posix-Lexing /theories/posix-lexing/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/posix-lexing/ Possibilistic_Noninterference /theories/possibilistic_noninterference/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/possibilistic_noninterference/ Power_Sum_Polynomials /theories/power_sum_polynomials/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/power_sum_polynomials/ Pratt_Certificate /theories/pratt_certificate/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/pratt_certificate/ Prefix_Free_Code_Combinators /theories/prefix_free_code_combinators/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/prefix_free_code_combinators/ Presburger-Automata /theories/presburger-automata/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/presburger-automata/ Prim_Dijkstra_Simple /theories/prim_dijkstra_simple/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/prim_dijkstra_simple/ Prime_Distribution_Elementary /theories/prime_distribution_elementary/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/prime_distribution_elementary/ Prime_Harmonic_Series /theories/prime_harmonic_series/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/prime_harmonic_series/ Prime_Number_Theorem /theories/prime_number_theorem/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/prime_number_theorem/ Priority_Queue_Braun /theories/priority_queue_braun/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/priority_queue_braun/ Priority_Search_Trees /theories/priority_search_trees/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/priority_search_trees/ Probabilistic_Noninterference /theories/probabilistic_noninterference/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/probabilistic_noninterference/ Probabilistic_Prime_Tests /theories/probabilistic_prime_tests/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/probabilistic_prime_tests/ Probabilistic_System_Zoo /theories/probabilistic_system_zoo/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/probabilistic_system_zoo/ Probabilistic_Timed_Automata /theories/probabilistic_timed_automata/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/probabilistic_timed_automata/ Probabilistic_While /theories/probabilistic_while/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/probabilistic_while/ Program-Conflict-Analysis /theories/program-conflict-analysis/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/program-conflict-analysis/ Progress_Tracking /theories/progress_tracking/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/progress_tracking/ Projective_Geometry /theories/projective_geometry/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/projective_geometry/ Projective_Measurements /theories/projective_measurements/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/projective_measurements/ Promela /theories/promela/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/promela/ Proof_Strategy_Language /theories/proof_strategy_language/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/proof_strategy_language/ Propositional_Proof_Systems /theories/propositional_proof_systems/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/propositional_proof_systems/ PropResPI /theories/proprespi/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/proprespi/ Prpu_Maxflow /theories/prpu_maxflow/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/prpu_maxflow/ PSemigroupsConvolution /theories/psemigroupsconvolution/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/psemigroupsconvolution/ PseudoHoops /theories/pseudohoops/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/pseudohoops/ Psi_Calculi /theories/psi_calculi/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/psi_calculi/ Ptolemys_Theorem /theories/ptolemys_theorem/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ptolemys_theorem/ Public_Announcement_Logic /theories/public_announcement_logic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/public_announcement_logic/ QHLProver /theories/qhlprover/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/qhlprover/ QR_Decomposition /theories/qr_decomposition/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/qr_decomposition/ Quantales /theories/quantales/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/quantales/ Quasi_Borel_Spaces /theories/quasi_borel_spaces/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/quasi_borel_spaces/ Quaternions /theories/quaternions/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/quaternions/ Query_Optimization /theories/query_optimization/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/query_optimization/ Quick_Sort_Cost /theories/quick_sort_cost/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/quick_sort_cost/ Ramsey-Infinite /theories/ramsey-infinite/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ramsey-infinite/ Random_BSTs /theories/random_bsts/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/random_bsts/ Random_Graph_Subgraph_Threshold /theories/random_graph_subgraph_threshold/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/random_graph_subgraph_threshold/ Randomised_BSTs /theories/randomised_bsts/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/randomised_bsts/ Randomised_Social_Choice /theories/randomised_social_choice/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/randomised_social_choice/ Rank_Nullity_Theorem /theories/rank_nullity_theorem/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/rank_nullity_theorem/ Real_Impl /theories/real_impl/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/real_impl/ Real_Power /theories/real_power/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/real_power/ Real_Time_Deque /theories/real_time_deque/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/real_time_deque/ Recursion-Addition /theories/recursion-addition/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/recursion-addition/ Recursion-Theory-I /theories/recursion-theory-i/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/recursion-theory-i/ Refine_Imperative_HOL /theories/refine_imperative_hol/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/refine_imperative_hol/ Refine_Monadic /theories/refine_monadic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/refine_monadic/ RefinementReactive /theories/refinementreactive/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/refinementreactive/ Regex_Equivalence /theories/regex_equivalence/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/regex_equivalence/ Registers /theories/registers/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/registers/ Regression_Test_Selection /theories/regression_test_selection/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/regression_test_selection/ Regular-Sets /theories/regular-sets/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/regular-sets/ Regular_Algebras /theories/regular_algebras/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/regular_algebras/ Regular_Tree_Relations /theories/regular_tree_relations/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/regular_tree_relations/ Relation_Algebra /theories/relation_algebra/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/relation_algebra/ Relational-Incorrectness-Logic /theories/relational-incorrectness-logic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/relational-incorrectness-logic/ Relational_Disjoint_Set_Forests /theories/relational_disjoint_set_forests/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/relational_disjoint_set_forests/ Relational_Forests /theories/relational_forests/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/relational_forests/ Relational_Method /theories/relational_method/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/relational_method/ Relational_Minimum_Spanning_Trees /theories/relational_minimum_spanning_trees/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/relational_minimum_spanning_trees/ Relational_Paths /theories/relational_paths/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/relational_paths/ Rep_Fin_Groups /theories/rep_fin_groups/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/rep_fin_groups/ Residuated_Lattices /theories/residuated_lattices/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/residuated_lattices/ ResiduatedTransitionSystem /theories/residuatedtransitionsystem/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/residuatedtransitionsystem/ Resolution_FOL /theories/resolution_fol/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/resolution_fol/ Rewrite_Properties_Reduction /theories/rewrite_properties_reduction/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/rewrite_properties_reduction/ Rewriting_Z /theories/rewriting_z/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/rewriting_z/ Ribbon_Proofs /theories/ribbon_proofs/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ribbon_proofs/ RIPEMD-160-SPARK /theories/ripemd-160-spark/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/ripemd-160-spark/ Risk_Free_Lending /theories/risk_free_lending/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/risk_free_lending/ Robbins-Conjecture /theories/robbins-conjecture/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/robbins-conjecture/ ROBDD /theories/robdd/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/robdd/ Robinson_Arithmetic /theories/robinson_arithmetic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/robinson_arithmetic/ Root_Balanced_Tree /theories/root_balanced_tree/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/root_balanced_tree/ Roth_Arithmetic_Progressions /theories/roth_arithmetic_progressions/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/roth_arithmetic_progressions/ Routing /theories/routing/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/routing/ Roy_Floyd_Warshall /theories/roy_floyd_warshall/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/roy_floyd_warshall/ RSAPSS /theories/rsapss/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/rsapss/ Safe_Distance /theories/safe_distance/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/safe_distance/ Safe_OCL /theories/safe_ocl/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/safe_ocl/ Safe_Range_RC /theories/safe_range_rc/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/safe_range_rc/ SATSolverVerification /theories/satsolververification/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/satsolververification/ Saturation_Framework /theories/saturation_framework/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/saturation_framework/ Saturation_Framework_Extensions /theories/saturation_framework_extensions/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/saturation_framework_extensions/ SC_DOM_Components /theories/sc_dom_components/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sc_dom_components/ SCC_Bloemen_Sequential /theories/scc_bloemen_sequential/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/scc_bloemen_sequential/ Schutz_Spacetime /theories/schutz_spacetime/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/schutz_spacetime/ SDS_Impossibility /theories/sds_impossibility/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sds_impossibility/ Search the Archive /search/ Mon, 01 Jan 0001 00:00:00 +0000 /search/ Secondary_Sylow /theories/secondary_sylow/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/secondary_sylow/ Security_Protocol_Refinement /theories/security_protocol_refinement/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/security_protocol_refinement/ Selection_Heap_Sort /theories/selection_heap_sort/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/selection_heap_sort/ SenSocialChoice /theories/sensocialchoice/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sensocialchoice/ Separata /theories/separata/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/separata/ Separation_Algebra /theories/separation_algebra/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/separation_algebra/ Separation_Logic_Imperative_HOL /theories/separation_logic_imperative_hol/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/separation_logic_imperative_hol/ Separation_Logic_Unbounded /theories/separation_logic_unbounded/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/separation_logic_unbounded/ Sepref_Basic /theories/sepref_basic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sepref_basic/ Sepref_IICF /theories/sepref_iicf/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sepref_iicf/ Sepref_Prereq /theories/sepref_prereq/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sepref_prereq/ SequentInvertibility /theories/sequentinvertibility/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sequentinvertibility/ Shadow_DOM /theories/shadow_dom/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/shadow_dom/ Shadow_SC_DOM /theories/shadow_sc_dom/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/shadow_sc_dom/ Shivers-CFA /theories/shivers-cfa/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/shivers-cfa/ ShortestPath /theories/shortestpath/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/shortestpath/ Show /theories/show/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/show/ SIFPL /theories/sifpl/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sifpl/ SIFUM_Type_Systems /theories/sifum_type_systems/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sifum_type_systems/ Sigma_Commit_Crypto /theories/sigma_commit_crypto/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sigma_commit_crypto/ Signature_Groebner /theories/signature_groebner/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/signature_groebner/ Simpl /theories/simpl/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/simpl/ Simple_Firewall /theories/simple_firewall/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/simple_firewall/ Simplex /theories/simplex/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/simplex/ Simplicial_complexes_and_boolean_functions /theories/simplicial_complexes_and_boolean_functions/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/simplicial_complexes_and_boolean_functions/ SimplifiedOntologicalArgument /theories/simplifiedontologicalargument/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/simplifiedontologicalargument/ Skew_Heap /theories/skew_heap/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/skew_heap/ Skip_Lists /theories/skip_lists/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/skip_lists/ Slicing /theories/slicing/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/slicing/ Sliding_Window_Algorithm /theories/sliding_window_algorithm/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sliding_window_algorithm/ SM /theories/sm/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sm/ SM_Base /theories/sm_base/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sm_base/ Smith_Normal_Form /theories/smith_normal_form/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/smith_normal_form/ Smooth_Manifolds /theories/smooth_manifolds/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/smooth_manifolds/ Solidity /theories/solidity/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/solidity/ Sophomores_Dream /theories/sophomores_dream/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sophomores_dream/ Sort_Encodings /theories/sort_encodings/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sort_encodings/ Source_Coding_Theorem /theories/source_coding_theorem/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/source_coding_theorem/ SPARCv8 /theories/sparcv8/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sparcv8/ SpecCheck /theories/speccheck/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/speccheck/ Special_Function_Bounds /theories/special_function_bounds/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/special_function_bounds/ Splay_Tree /theories/splay_tree/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/splay_tree/ Sqrt_Babylonian /theories/sqrt_babylonian/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sqrt_babylonian/ Stable_Matching /theories/stable_matching/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/stable_matching/ Stalnaker_Logic /theories/stalnaker_logic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/stalnaker_logic/ Statecharts /theories/statecharts/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/statecharts/ Stateful_Protocol_Composition_and_Typing /theories/stateful_protocol_composition_and_typing/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/stateful_protocol_composition_and_typing/ Statistics /statistics/ Mon, 01 Jan 0001 00:00:00 +0000 /statistics/ - 712 Entries 436 Authors ~226,200 Lemmas ~3,664,200 Lines of Code Most used AFP entries: Name Used by ? entries 1. List-Index 22 2. Collections 19 3. Show 16 4. Deriving 13 5. Coinductive 12 6. Jordan_Normal_Form 12 7. Polynomial_Factorization 12 8. Regular-Sets 12 9. + 712 Entries 436 Authors ~226,200 Lemmas ~3,664,100 Lines of Code Most used AFP entries: Name Used by ? entries 1. List-Index 22 2. Collections 19 3. Show 16 4. Deriving 13 5. Coinductive 12 6. Jordan_Normal_Form 12 7. Polynomial_Factorization 12 8. Regular-Sets 12 9. Stellar_Quorums /theories/stellar_quorums/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/stellar_quorums/ Stern_Brocot /theories/stern_brocot/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/stern_brocot/ Stewart_Apollonius /theories/stewart_apollonius/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/stewart_apollonius/ Stirling_Formula /theories/stirling_formula/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/stirling_formula/ Stochastic_Matrices /theories/stochastic_matrices/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/stochastic_matrices/ Stone_Algebras /theories/stone_algebras/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/stone_algebras/ Stone_Kleene_Relation_Algebras /theories/stone_kleene_relation_algebras/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/stone_kleene_relation_algebras/ Stone_Relation_Algebras /theories/stone_relation_algebras/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/stone_relation_algebras/ Store_Buffer_Reduction /theories/store_buffer_reduction/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/store_buffer_reduction/ Stream-Fusion /theories/stream-fusion/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/stream-fusion/ Stream_Fusion_Code /theories/stream_fusion_code/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/stream_fusion_code/ Strong_Security /theories/strong_security/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/strong_security/ Sturm_Sequences /theories/sturm_sequences/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sturm_sequences/ Sturm_Tarski /theories/sturm_tarski/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sturm_tarski/ Stuttering_Equivalence /theories/stuttering_equivalence/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/stuttering_equivalence/ Submission /webapp/submission/ Mon, 01 Jan 0001 00:00:00 +0000 /webapp/submission/ { if (Number.isInteger(event.data)) { this.style.height=event.data + 32 + "px" } }) }).call(this)' const base = 'https:\/\/ci.isabelle.systems\/afp-submission' document.getElementById("iframe").src = base + '/submission' + window.location.search Submission /webapp/submit/ Mon, 01 Jan 0001 00:00:00 +0000 /webapp/submit/ { if (Number.isInteger(event.data)) { this.style.height=event.data + 32 + "px" } }) }).call(this)' const base = 'https:\/\/ci.isabelle.systems\/afp-submission' document.getElementById("iframe").src = base + '/submit' + window.location.search Subresultants /theories/subresultants/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/subresultants/ Subset_Boolean_Algebras /theories/subset_boolean_algebras/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/subset_boolean_algebras/ SumSquares /theories/sumsquares/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sumsquares/ Sunflowers /theories/sunflowers/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/sunflowers/ SuperCalc /theories/supercalc/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/supercalc/ Surprise_Paradox /theories/surprise_paradox/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/surprise_paradox/ Symmetric_Polynomials /theories/symmetric_polynomials/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/symmetric_polynomials/ Syntax_Independent_Logic /theories/syntax_independent_logic/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/syntax_independent_logic/ Szemeredi_Regularity /theories/szemeredi_regularity/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/szemeredi_regularity/ Szpilrajn /theories/szpilrajn/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/szpilrajn/ Tail_Recursive_Functions /theories/tail_recursive_functions/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/tail_recursive_functions/ Tarskis_Geometry /theories/tarskis_geometry/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/tarskis_geometry/ Taylor_Models /theories/taylor_models/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/taylor_models/ TESL_Language /theories/tesl_language/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/tesl_language/ Three_Circles /theories/three_circles/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/three_circles/ Timed_Automata /theories/timed_automata/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/timed_automata/ TLA /theories/tla/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/tla/ Topological_Semantics /theories/topological_semantics/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/topological_semantics/ Topology /theories/topology/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/topology/ TortoiseHare /theories/tortoisehare/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/tortoisehare/ Transcendence_Series_Hancl_Rucki /theories/transcendence_series_hancl_rucki/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/transcendence_series_hancl_rucki/ Transformer_Semantics /theories/transformer_semantics/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/transformer_semantics/ Transition_Systems_and_Automata /theories/transition_systems_and_automata/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/transition_systems_and_automata/ Transitive-Closure /theories/transitive-closure/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/transitive-closure/ Transitive-Closure-II /theories/transitive-closure-ii/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/transitive-closure-ii/ Transitive_Models /theories/transitive_models/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/transitive_models/ Treaps /theories/treaps/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/treaps/ Tree-Automata /theories/tree-automata/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/tree-automata/ Tree_Decomposition /theories/tree_decomposition/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/tree_decomposition/ Triangle /theories/triangle/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/triangle/ Trie /theories/trie/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/trie/ Twelvefold_Way /theories/twelvefold_way/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/twelvefold_way/ Tycon /theories/tycon/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/tycon/ Types_Tableaus_and_Goedels_God /theories/types_tableaus_and_goedels_god/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/types_tableaus_and_goedels_god/ Types_To_Sets_Extension /theories/types_to_sets_extension/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/types_to_sets_extension/ Undirected_Graph_Theory /theories/undirected_graph_theory/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/undirected_graph_theory/ Universal_Hash_Families /theories/universal_hash_families/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/universal_hash_families/ Universal_Turing_Machine /theories/universal_turing_machine/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/universal_turing_machine/ UpDown_Scheme /theories/updown_scheme/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/updown_scheme/ UPF /theories/upf/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/upf/ UPF_Firewall /theories/upf_firewall/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/upf_firewall/ UTP /theories/utp/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/utp/ UTP-Toolkit /theories/utp-toolkit/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/utp-toolkit/ Valuation /theories/valuation/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/valuation/ Van_der_Waerden /theories/van_der_waerden/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/van_der_waerden/ Van_Emde_Boas_Trees /theories/van_emde_boas_trees/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/van_emde_boas_trees/ VectorSpace /theories/vectorspace/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/vectorspace/ VeriComp /theories/vericomp/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/vericomp/ Verified-Prover /theories/verified-prover/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/verified-prover/ Verified_SAT_Based_AI_Planning /theories/verified_sat_based_ai_planning/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/verified_sat_based_ai_planning/ VerifyThis2018 /theories/verifythis2018/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/verifythis2018/ VerifyThis2019 /theories/verifythis2019/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/verifythis2019/ Vickrey_Clarke_Groves /theories/vickrey_clarke_groves/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/vickrey_clarke_groves/ Virtual_Substitution /theories/virtual_substitution/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/virtual_substitution/ VolpanoSmith /theories/volpanosmith/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/volpanosmith/ VYDRA_MDL /theories/vydra_mdl/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/vydra_mdl/ WebAssembly /theories/webassembly/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/webassembly/ Weight_Balanced_Trees /theories/weight_balanced_trees/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/weight_balanced_trees/ Weighted_Arithmetic_Geometric_Mean /theories/weighted_arithmetic_geometric_mean/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/weighted_arithmetic_geometric_mean/ Weighted_Path_Order /theories/weighted_path_order/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/weighted_path_order/ Well_Quasi_Orders /theories/well_quasi_orders/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/well_quasi_orders/ Wetzels_Problem /theories/wetzels_problem/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/wetzels_problem/ WHATandWHERE_Security /theories/whatandwhere_security/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/whatandwhere_security/ Winding_Number_Eval /theories/winding_number_eval/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/winding_number_eval/ WOOT_Strong_Eventual_Consistency /theories/woot_strong_eventual_consistency/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/woot_strong_eventual_consistency/ Word_Lib /theories/word_lib/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/word_lib/ WorkerWrapper /theories/workerwrapper/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/workerwrapper/ X86_Semantics /theories/x86_semantics/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/x86_semantics/ XML /theories/xml/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/xml/ Youngs_Inequality /theories/youngs_inequality/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/youngs_inequality/ Zeta_3_Irrational /theories/zeta_3_irrational/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/zeta_3_irrational/ Zeta_Function /theories/zeta_function/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/zeta_function/ ZFC_in_HOL /theories/zfc_in_hol/ Mon, 01 Jan 0001 00:00:00 +0000 /theories/zfc_in_hol/ diff --git a/web/statistics/index.html b/web/statistics/index.html --- a/web/statistics/index.html +++ b/web/statistics/index.html @@ -1,337 +1,337 @@ Statistics - Archive of Formal Proofs - + - +

Statistics

- +
712 Entries
436 Authors
~226,200 Lemmas
~3,664,200~3,664,100 Lines of Code

Most used AFP entries:

Name Used by ? entries
1. List-Index 22
2. Collections 19
3. Show 16
4. Deriving 13
5. Coinductive 12
6. Jordan_Normal_Form 12
7. Polynomial_Factorization 12
8. Regular-Sets 12
9. Landau_Symbols 11
10. Abstract-Rewriting 10
11. Automatic_Refinement 10
12. Native_Word 10

Growth in number of entries:

Growth in lines of code:

Growth in number of authors:

Size of entries:

\ No newline at end of file