--- a/Exercises/ROOT Sat Jul 29 19:34:17 2017 +0200
+++ b/Exercises/ROOT Sat Jul 29 19:34:29 2017 +0200
@@ -84,6 +84,19 @@
theories ex13
document_files "root.tex" "exercise.sty" "build"
+session "ex14_base" in "ex14" = HOL +
+ theories
+ "~~/src/HOL/Library/Multiset"
+ "~~/src/HOL/Library/Code_Target_Nat"
+ "~~/src/HOL/Library/RBT"
+ "~~/src/HOL/Library/Char_ord"
+ "~~/src/HOL/Library/Code_Char"
+
+session "ex14" in "ex14" = ex14_base +
+ options [document = pdf, document_output = "generated", document_variants = "ex14"]
+ theories ex14
+ document_files "root.tex" "exercise.sty" "build"
+
session exam_basis = HOL +
theories [document=false]
"~~/src/HOL/Library/Multiset"
@@ -94,3 +107,32 @@
options [document = pdf, document_output = "generated", document_variants = "exam"]
theories[show_question_marks=false] exam
document_files "root.tex" "exercise.sty" "build"
+
+session hwsol_basis = HOL +
+ theories
+ "~~/src/HOL/Library/Multiset"
+ "../../Public/Demos/BST_Demo"
+ "~~/src/HOL/Library/Tree"
+ "~~/src/HOL/Data_Structures/Tree23_Set"
+ "../../Public/Thys/Trie1"
+ "../../Public/Thys/Leftist_Heap"
+ "../../Public/Thys/Heap_Prelim"
+ "../../Public/Thys/Skew_Heap"
+
+session "hwsol" in "hwsol" = hwsol_basis +
+ options [document = pdf, document_output = "generated", document_variants = "hwsol"]
+ theories[show_question_marks=false]
+ hwsol01
+ hwsol02
+ hwsol03
+ hwsol04
+ hwsol05
+ hwsol06
+ hwsol07
+ hwsol08
+ hwsol09
+ hwsol10
+ hwsol11
+ hwsol12
+ hwsol13
+ document_files "root.tex" "build"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/ex13/sol13.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,213 @@
+(*<*)
+theory sol13
+imports "../../../Public/Thys/Skew_Heap"
+begin
+(*>*)
+
+text {* \ExerciseSheet{13}{21.~7.~2017} *}
+
+text \<open>{
+ \large\bf Presentation of Mini-Projects}
+
+ You are invited, on a voluntary basis, to give a short presentation
+ of your mini-projects in the tutorial on July 28.
+
+ Depending on how many presentations we have, the time slots
+ will be 5 to 10 minutes, plus 2 minutes for questions.
+
+ If you are interested, please write me a short email until Thursday, July 27.
+\<close>
+
+(* FIXME: Should there be a hint about function "abs" for the use in \<Phi>? *)
+
+text \<open>
+\Exercise{Double-Ended Queues}
+
+Design a double-ended queue where all operations have constant-time amortized
+complexity. Prove functional correctness and constant-time amortized complexity.
+
+For your proofs, it is enough to count the number of newly allocated list cells.
+You may assume that operations @{term \<open>rev xs\<close>} and @{term \<open>xs@ys\<close>} allocate \<open>O(|xs|)\<close>
+cells.
+
+Explanation: A double-ended queue is like a queue with two further operations:
+Function \<open>enq_front\<close> adds an element at the front (whereas \<open>enq\<close> adds an element at the back).
+Function \<open>deq_back\<close> removes an element at the back (whereas \<open>deq\<close> removes an element at the front).
+Here is a formal specification where the double ended queue is just a list:
+\<close>
+
+abbreviation (input) enq_list :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+"enq_list x xs \<equiv> xs @ [x]"
+
+abbreviation (input) enq_front_list :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+"enq_front_list x xs \<equiv> x # xs"
+
+abbreviation (input) deq_list :: "'a list \<Rightarrow> 'a list" where
+"deq_list xs \<equiv> tl xs"
+
+abbreviation (input) deq_back_list :: "'a list \<Rightarrow> 'a list" where
+"deq_back_list xs \<equiv> butlast xs"
+
+text\<open>
+Hint: You may want to start with the \<open>Queue\<close> implementation in \<open>Thys/Amortized_Examples\<close>.\<close>
+
+(*<*)
+
+type_synonym 'a queue = "'a list * 'a list"
+
+fun list_of :: "'a queue \<Rightarrow> 'a list" where
+"list_of(xs,ys) = ys @ rev xs"
+
+(*>*)
+
+(*<*)
+
+(* Definitions *)
+
+definition init :: "'a queue" where
+"init = ([],[])"
+
+fun enq :: "'a \<Rightarrow> 'a queue \<Rightarrow> 'a queue" where
+"enq x (xs,ys) = (x#xs, ys)"
+
+fun enq_front :: "'a \<Rightarrow> 'a queue \<Rightarrow> 'a queue" where
+"enq_front x (xs,ys) = (xs, x#ys)"
+
+fun deq :: "'a queue \<Rightarrow> 'a queue" where
+"deq (xs,ys) =
+ (if ys = []
+ then let n = length xs div 2
+ in (take n xs, tl(rev (drop n xs)))
+ else (xs, tl ys))"
+
+fun deq_back :: "'a queue \<Rightarrow> 'a queue" where
+"deq_back (xs,ys) =
+ (if xs = []
+ then let n = length ys div 2
+ in (tl(rev (drop n ys)), take n ys)
+ else (tl xs, ys))"
+
+(*>*)
+(* Functional Correctness *)
+
+lemma "list_of init = []"
+(*<*)
+ by (auto simp: init_def)
+(*>*)
+
+lemma "list_of(enq x q) = enq_list x (list_of q)"
+(*<*)
+by(cases q) auto
+(*>*)
+
+lemma "list_of(enq_front x q) = enq_front_list x (list_of q)"
+(*<*)
+by(cases q) auto
+(*>*)
+
+lemma "list_of q \<noteq> [] \<Longrightarrow> list_of(deq q) = deq_list (list_of q)"
+(*<*)
+proof (cases q)
+ case [simp]: (Pair a b)
+
+ assume "list_of q \<noteq> []"
+ hence "\<not> length a \<le> length a div 2" and [simp]: "a\<noteq>[]" if "b=[]" using that
+ by (cases a; auto)+
+
+ then show ?thesis
+ by (auto simp: Let_def rev_drop rev_take tl_append2[symmetric]
+ simp del: tl_append2)
+
+qed
+(*>*)
+
+(*<*)
+lemma rev_tl_butlast_rev: "rev (tl xs) = butlast(rev xs)"
+by (metis One_nat_def butlast_conv_take drop_0 drop_Suc drop_rev rev_swap)
+
+lemma butlast_append2: "b\<noteq>[] \<Longrightarrow> butlast (a@b) = a@butlast b"
+ by (auto simp: butlast_append)
+
+lemma [simp]: "b \<noteq> [] \<Longrightarrow> \<not> length b \<le> length b div 2" by (cases b) auto
+(*>*)
+
+lemma "list_of q \<noteq> [] \<Longrightarrow> list_of(deq_back q) = deq_back_list (list_of q)"
+(*<*)
+ apply (cases q)
+ apply (auto simp: Let_def rev_tl_butlast_rev butlast_append2[symmetric])
+ done
+(*>*)
+
+
+(*<*)
+
+fun t_enq :: "'a \<Rightarrow> 'a queue \<Rightarrow> int" where
+(*<*)
+"t_enq x (xs,ys) = 1"
+(*>*)
+
+fun t_enq_front :: "'a \<Rightarrow> 'a queue \<Rightarrow> int" where
+(*<*)
+"t_enq_front x (xs,ys) = 1"
+(*>*)
+
+
+fun t_deq :: "'a queue \<Rightarrow> int" where
+(*<*)
+"t_deq (xs,ys) =
+ (if ys = []
+ then let n = length xs div 2
+ in int(n + (length xs - n))
+ else 0)"
+(*>*)
+
+fun t_deq_back :: "'a queue \<Rightarrow> int" where
+(*<*)
+"t_deq_back (xs,ys) =
+ (if xs = []
+ then let n = length ys; m = n div 2
+ in int((length ys - n) + n)
+ else 0)"
+(*>*)
+
+fun \<Phi> :: "'a queue \<Rightarrow> int"
+(*<*)
+ where
+"\<Phi> (xs,ys) = abs(int(length xs) - int(length ys))"
+(*>*)
+
+lemma "\<Phi> q \<ge> 0"
+(*<*)
+by(cases q) simp
+(*>*)
+
+lemma "\<Phi> init = 0"
+(*<*)
+by(simp add: init_def)
+(*>*)
+
+lemma "t_enq x q + \<Phi>(enq x q) - \<Phi> q \<le> 2"
+(*<*)
+by(cases q) auto
+(*>*)
+
+lemma "t_enq_front x q + \<Phi>(enq_front x q) - \<Phi> q \<le> 2"
+(*<*)
+by(cases q) auto
+(*>*)
+
+lemma "t_deq q + \<Phi>(deq q) - \<Phi> q \<le> 2"
+(*<*)
+by(cases q) (auto simp: Let_def)
+(*>*)
+
+lemma "t_deq_back q + \<Phi>(deq_back q) - \<Phi> q \<le> 2"
+(*<*)
+by(cases q) (auto simp: Let_def)
+(*>*)
+
+(*>*)
+
+(*<*)
+end
+(*>*)
Binary file Exercises/ex14.pdf has changed
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/ex14/document/build Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,8 @@
+#!/bin/sh
+
+${ISABELLE_TOOL} latex -o sty root.tex && \
+${ISABELLE_TOOL} latex -o pdf root.tex && \
+${ISABELLE_TOOL} latex -o pdf root.tex && \
+${ISABELLE_TOOL} latex -o pdf root.tex && \
+cp root.pdf ../../../$2.pdf
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/ex14/document/exercise.sty Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,51 @@
+
+\newcommand{\Lecture}{Functional Data Structures}
+\newcommand{\Semester}{SS 2017}
+\newcommand{\Prof}{Prof.~Tobias~Nipkow,~Ph.D.}
+\newcommand{\Tutor}{Dr. Peter Lammich}
+
+\newcounter{sheet}
+\newcounter{homework}
+
+\newlength{\abslength}\setlength{\abslength}{2mm plus 1mm minus 1mm}
+\newcommand{\abs}{\par\vspace{\abslength}}
+
+\newlength{\TUMBr}\settowidth{\TUMBr}{{\bf Technische Universit{\"a}t M{\"u}nchen}}
+\newcommand{\Header}[5]{{\bf
+ \makebox[\TUMBr]{Technische Universit{\"a}t M{\"u}nchen} \hfill #3\\
+ \makebox[\TUMBr]{Institut f{\"u}r Informatik} \hfill #4\\
+ \makebox[\TUMBr]{#1} \hfill #5\\
+ \makebox[\TUMBr]{#2}}\abs}
+
+\newcommand{\Title}[1]{%
+ \begin{center}{\LARGE\bf\Lecture}\\[1ex]{\bf Exercise Sheet #1}\end{center}}
+
+\newcommand{\ExerciseSheet}[2]{%
+ \pagestyle{empty}%
+ \setcounter{sheet}{#1}%
+ \vspace*{-2cm}\Header{\Prof}{\Tutor}{\Semester}{#2}{}\vspace*{1cm}%
+ \Title{#1}\abs}
+
+\newcounter{exercise}
+\newcommand{\Exercise}[1]{%
+ \refstepcounter{exercise}%
+ \pagebreak[3]%
+ \relax%
+ \vspace{0.8em}%
+ \subsection*{{Exercise \arabic{sheet}.\arabic{exercise}\ \ \normalfont\sffamily #1}}}
+
+\newcommand{\Homework}[2]{%
+ \pagebreak[3]%
+ \relax%
+ \vspace{0.8em}%
+ \subsection*{{Homework \arabic{sheet}\ \ \normalfont\sffamily #1}}%
+ \emph{Submission until Friday, #2, 11:59am.}}
+
+\newcommand{\NumHomework}[2]{%
+ \refstepcounter{homework}%
+ \pagebreak[3]%
+ \relax%
+ \vspace{0.8em}%
+ \subsection*{{Homework \arabic{sheet}.\arabic{homework}\ \ \normalfont\sffamily #1}}%
+ \emph{Submission until Friday, #2, 11:59am.}}
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/ex14/document/root.tex Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,66 @@
+\documentclass[11pt,a4paper]{scrartcl}
+\usepackage{isabelle,isabellesym}
+\usepackage{exercise}
+\usepackage{amsmath}
+
+% further packages required for unusual symbols (see also
+% isabellesym.sty), use only when needed
+
+\usepackage{amssymb}
+ %for \<leadsto>, \<box>, \<diamond>, \<sqsupset>, \<mho>, \<Join>,
+ %\<lhd>, \<lesssim>, \<greatersim>, \<lessapprox>, \<greaterapprox>,
+ %\<triangleq>, \<yen>, \<lozenge>
+
+%\usepackage[greek,english]{babel}
+ %option greek for \<euro>
+ %option english (default language) for \<guillemotleft>, \<guillemotright>
+
+\usepackage[latin1]{inputenc}
+ %for \<onesuperior>, \<onequarter>, \<twosuperior>, \<onehalf>,
+ %\<threesuperior>, \<threequarters>, \<degree>
+
+\usepackage[only,bigsqcap]{stmaryrd}
+ %for \<Sqinter>
+
+%\usepackage{eufrak}
+ %for \<AA> ... \<ZZ>, \<aa> ... \<zz> (also included in amssymb)
+
+%\usepackage{textcomp}
+ %for \<cent>, \<currency>
+
+% this should be the last package used
+\usepackage{pdfsetup}
+
+% urls in roman style, theory text in math-similar italics
+\urlstyle{rm}
+\isabellestyle{it}
+
+% for uniform font size
+%\renewcommand{\isastyle}{\isastyleminor}
+
+\begin{document}
+
+% sane default for proof documents
+\parindent 0pt\parskip 0.5ex
+
+
+\renewcommand{\isachardoublequote}{`\"}
+\renewcommand{\isachardoublequoteopen}{``}
+\renewcommand{\isachardoublequoteclose}{''}
+\renewcommand{\isacharunderscore}{\_}
+\renewcommand{\isacharunderscorekeyword}{\_}
+
+
+% generated text of all theories
+\input{session}
+
+% optional bibliography
+%\bibliographystyle{abbrv}
+%\bibliography{root}
+
+\end{document}
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
--- a/Exercises/ex14/ex14.thy Sat Jul 29 19:34:17 2017 +0200
+++ b/Exercises/ex14/ex14.thy Sat Jul 29 19:34:29 2017 +0200
@@ -1,3 +1,4 @@
+text {* \ExerciseSheet{14}{28.~7.~2017} *}
theory ex14
imports
"~~/src/HOL/Library/Multiset"
@@ -7,92 +8,235 @@
"~~/src/HOL/Library/Code_Char"
begin
- definition "freq ws \<equiv> count (mset ws)"
- definition "incr1 k m \<equiv> case m k of None \<Rightarrow> m(k\<mapsto>1) | Some v \<Rightarrow> m(k\<mapsto>Suc v)"
-
- definition "freq1 ws \<equiv> fold incr1 ws Map.empty"
-
- definition "abs_freq1 m k \<equiv> case m k of None \<Rightarrow> 0 | Some v \<Rightarrow> v"
-
- lemma [simp]: "abs_freq1 Map.empty = (\<lambda>_. 0)"
- unfolding abs_freq1_def by auto
+ text \<open>
+ \Exercise{Word Frequency --- Down to ML code}
+
+ Your task is to develop a program that reads a corpus and
+ prints the words in the corpus together with their frequencies,
+ sorted by descending frequencies.
+
+ Except input and output, your program shall be formalized in Isabelle/HOL.
+ A corpus is described as @{typ \<open>'a list\<close>},
+ and the result is described by @{typ \<open>('a \<times> nat) list\<close>}
+ \<close>
+
+ text \<open>The frequency of a word in a corpus can be specified as:\<close>
+ definition freq :: "'a list \<Rightarrow> 'a \<Rightarrow> nat"
+ where "freq ws = count (mset ws)"
- lemma [simp]: "abs_freq1 (freq1 ws) = freq ws"
- proof
- fix k
- have "abs_freq1 (fold incr1 ws m) k = freq ws k + abs_freq1 m k" for m k
- apply (induction ws arbitrary: m)
- apply (auto simp: freq_def abs_freq1_def incr1_def split: option.splits)
- done
- from this[of Map.empty k] show "abs_freq1 (freq1 ws) k = freq ws k"
- by (auto simp: freq1_def)
- qed
-
+ text \<open>
+ Specify a predicate that characterizes a correct result.
+ Note: If words have the same frequency, any order is allowed.
+ \<close>
+ definition is_freq_list :: "'a list \<Rightarrow> ('a \<times> nat) list \<Rightarrow> bool"
+ (*<*)
+ where
+ "is_freq_list ws fl \<equiv> (
+ distinct (map fst fl)
+ \<and> (sorted (rev (map snd fl)))
+ \<and> (\<forall>w f. (w,f)\<in>set fl \<longleftrightarrow> f=freq ws w \<and> f>0)
+ )"
+ (*>*)
+
+ text \<open>Tests:\<close>
+ lemma \<open>is_freq_list [''a'',''b'',''b'',''c'',''c''] [(''c'',2),(''b'',2),(''a'',1)]\<close>
+ (*<*)
+ unfolding is_freq_list_def freq_def by auto
+ (*>*)
+ lemma \<open>is_freq_list [''a'',''b'',''b'',''c'',''c''] [(''b'',2),(''c'',2),(''a'',1)]\<close>
+ (*<*)
+ unfolding is_freq_list_def freq_def by auto
+ (*>*)
+ lemma \<open>\<not>is_freq_list [''a'',''b'',''b'',''c'',''c''] [(''b'',2),(''c'',3),(''a'',1)]\<close>
+ (*<*)
+ unfolding is_freq_list_def freq_def by auto
+ (*>*)
+ lemma \<open>\<not>is_freq_list [''a'',''b'',''b'',''c'',''c''] [(''a'',1),(''c'',2),(''b'',2)]\<close>
+ (*<*)
+ unfolding is_freq_list_def freq_def by auto
+ (*>*)
+ lemma \<open>\<not>is_freq_list [''a'',''b'',''b'',''c'',''c''] [(''b'',2),(''c'',2),(''b'',2),(''a'',1)]\<close>
+ (*<*)
+ unfolding is_freq_list_def freq_def by auto
+ (*>*)
+
+
+ section \<open>Refinement \#1\<close>
+ (* Refinement #1: Fold over word list and maintain current frequency *)
+ text \<open>Compute a function from words to their frequency by folding
+ over the corpus, starting with @{term \<open>\<lambda>_. 0::nat\<close>}. \<close>
+
+ definition incr1 :: "'a \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> 'a \<Rightarrow> nat"
+ (*<*)
+ where
+ "incr1 k f \<equiv> f(k := Suc (f k))"
+ (*>*)
+ definition "freq1 ws \<equiv> fold incr1 ws (\<lambda>_. 0)"
- definition "incr2 k m \<equiv> case RBT.lookup m k of None \<Rightarrow> RBT.insert k 1 m | Some v \<Rightarrow> RBT.insert k (Suc v) m"
+ lemma freq1_correct[simp]: "freq1 ws = freq ws"
+ (*<*)
+ proof -
+ have "fold incr1 ws f k = freq ws k + f k" for f k
+ by (induction ws arbitrary: f) (auto simp: freq_def incr1_def)
+ thus ?thesis unfolding freq1_def by auto
+ qed
+ (*>*)
- lemma incr2_abs[simp]: "RBT.lookup (incr2 k m) = incr1 k (RBT.lookup m)"
- unfolding incr2_def incr1_def
- by (auto split: option.split)
+ section \<open>Refinement \#2\<close>
+ (* Refinement #2: Frequency map by RBT *)
+ text \<open>
+ Use red black trees to implement the mapping from words to frequencies.
+ Words that do not occur in the corpus must not be mapped!
+
+ Use the RBT implementation from \<open>HOL/Library/RBT\<close>!
+ It provides, e.g., @{const RBT.empty}, @{const RBT.lookup}, @{const RBT.insert}.
+ \<close>
+ definition abs_fm :: "('a::linorder,nat) rbt \<Rightarrow> 'a \<Rightarrow> nat" where
+ "abs_fm t k \<equiv> case RBT.lookup t k of None \<Rightarrow> 0 | Some v \<Rightarrow> v"
+ definition inv_fm :: "('a::linorder,nat) rbt \<Rightarrow> bool" where
+ "inv_fm t \<equiv> (0 \<notin> ran (RBT.lookup t))"
- definition "freq2 ws \<equiv> fold incr2 ws RBT.empty"
-
- lemma freq2_abs[simp]: "RBT.lookup (freq2 ws) = freq1 ws"
- proof -
- have "RBT.lookup (fold incr2 ws t) = fold incr1 ws (RBT.lookup t)" for t
- by (induction ws arbitrary: t) auto
- thus ?thesis unfolding freq2_def freq1_def by auto
- qed
+ lemma empty2_correct[simp]:
+ "abs_fm RBT.empty = (\<lambda>_. 0)" "inv_fm RBT.empty"
+ (*<*)
+ by (auto simp: abs_fm_def inv_fm_def)
+ (*>*)
- export_code freq2 in SML module_name Freq
-
-
- definition "fsort ws \<equiv> rev (sort_key (freq ws) (remdups ws))"
- definition "fsort2 ws \<equiv> let t = freq2 ws in rev (sort_key (abs_freq1 (RBT.lookup t)) (remdups ws))"
+ definition incr2 :: "'a::linorder \<Rightarrow> ('a, nat) rbt \<Rightarrow> ('a, nat) rbt"
+ (*<*)
+ where
+ "incr2 k t \<equiv> case RBT.lookup t k of
+ None \<Rightarrow> RBT.insert k 1 t
+ | Some v \<Rightarrow> RBT.insert k (Suc v) t"
+ (*>*)
+
+ lemma incr2_correct[simp]:
+ "inv_fm t \<Longrightarrow> abs_fm (incr2 k t) = incr1 k (abs_fm t)"
+ "inv_fm t \<Longrightarrow> inv_fm (incr2 k t)"
+ (*<*)
+ by (auto simp: inv_fm_def abs_fm_def incr2_def incr1_def ran_def split: option.splits)
+ (*>*)
-
- definition "rdfreq ws \<equiv> fold (\<lambda>x (m,l). (add_mset x m, if x\<in>#m then l else x#l)) ws ({#},[])"
-
- lemma rdfreq_correct:
- assumes "rdfreq ws = (m,l)"
- shows "m=mset ws" "distinct l" "set l = set ws"
+ text \<open> Now we have refined the operations, we can refine the algorithm that uses the operations:\<close>
+ definition "freq2 ws \<equiv> fold incr2 ws RBT.empty"
+
+ lemma freq2_correct[simp]: "abs_fm (freq2 ws) = freq1 ws" "inv_fm (freq2 ws)"
+ (*<*)
proof -
- {
- fix m\<^sub>0 l\<^sub>0
- assume "fold (\<lambda>x (m,l). (add_mset x m, if x\<in>#m then l else x#l)) ws (m\<^sub>0,l\<^sub>0) = (m,l)"
- and "set_mset m\<^sub>0 = set l\<^sub>0" "distinct l\<^sub>0"
- hence "m = m\<^sub>0 + mset ws \<and> distinct l \<and> set l = set l\<^sub>0 \<union> set ws"
- apply (induction ws arbitrary: m\<^sub>0 l\<^sub>0 m l)
- apply (fastforce split: if_splits)+
- done
- }
- from this[of "{#}" "[]"] show "m=mset ws" "distinct l" "set l = set ws"
- using assms by (auto simp: rdfreq_def)
+ have "abs_fm (fold incr2 ws t) = fold incr1 ws (abs_fm t) \<and> inv_fm (fold incr2 ws t)"
+ if "inv_fm t" for t
+ using that
+ by (induction ws arbitrary: t) auto
+ from this[of RBT.empty] show "abs_fm (freq2 ws) = freq1 ws" "inv_fm (freq2 ws)"
+ unfolding freq2_def freq1_def by auto
qed
+ (*>*)
+
+
+ (*<*) (* NOT PART OF EXERCISE *)
+ (* Simple abstract implementation *)
+ definition "fsort_simple ws \<equiv> rev (sort_key snd (map (\<lambda>w. (w,freq ws w)) (remdups ws)))"
+
+ lemma insort_key_map:
+ "insort_key f (g k) (map g l) = map g (insort_key (f o g) k l)"
+ by (induction l) auto
+
+ lemma sort_key_map: "sort_key f (map g l) = map g (sort_key (f o g) l)"
+ by (induction l) (auto simp: o_def insort_key_map)
+
+ lemma fsort_simple_is_freq_list: "is_freq_list ws (fsort_simple ws)"
+ unfolding is_freq_list_def fsort_simple_def
+ apply (safe; simp?; (auto; fail)?)
+ subgoal by (simp add: rev_map[symmetric] sort_key_map o_def)
+ subgoal by (simp add: rev_map)
+ subgoal by (auto simp: freq_def) []
+ subgoal by (auto simp: freq_def) []
+ done
+ (*>*)
- xxx, ctd here:
- still not accurate, as result after sort may be different !
-
-
+ subsection \<open>Extracting Result from RBT\<close>
+ text \<open>Extract the desired result
+ --- list of pairs of words and their frequencies, sorted by frequency ---
+ from the red black tree. Use @{const RBT.entries}.
+ \<close>
+
+ definition fsort :: "'a::linorder list \<Rightarrow> ('a \<times> nat) list"
+ (*<*)
+ where "fsort ws \<equiv>
+ let fr = freq2 ws in
+ rev (sort_key snd (RBT.entries fr))"
+ (*>*)
+
+ (*<*)
+ lemma distinct_map_insort[simp]:
+ "distinct (map f (insort_key g k l)) \<longleftrightarrow> distinct (map f (k#l))"
+ by (induction l arbitrary: k) (auto simp: set_insort_key)
- (* TODO: Use RBT.keys/RBT.entries instead of remdups! *)
+ lemma distinct_map_sort[simp]: "distinct (map f (sort_key g l)) \<longleftrightarrow> distinct (map f l)"
+ by (induction l) auto
+
+ (* For this proof, we have to derive the following from the freq2-correct lemma.
+ This is required as we work on red-black trees directly
+ *)
+ lemma freq2_eq_Some_iff[simp]: "RBT.lookup (freq2 ws) w = Some f \<longleftrightarrow> (f>0 \<and> f = freq ws w)"
+ proof -
+ note X = freq2_correct[of ws, unfolded abs_fm_def]
+ note Y = X(1)[THEN fun_cong, of w]
+ note Z = X(2)[unfolded inv_fm_def]
+ from Y Z show ?thesis
+ apply (auto split: option.splits simp: ran_def)
+ using neq0_conv by fastforce
+ qed
+ (*>*)
+
+ text \<open>Prove that your function is correct.
+ Hint: You will need to prove some auxiliary lemmas on standard list functions.
+ Use \<open>find_theorems\<close> to search for existing lemmas.
+ Hint: A lemma of the form @{text \<open>RBT.lookup (freq2 ws) w = Some f \<longleftrightarrow> \<dots>\<close>},
+ derived from @{thm [source] freq2_correct freq1_correct} may be useful!
+ \<close>
+ lemma fsort_correct: "is_freq_list ws (fsort ws)"
+ (*<*)
+ unfolding is_freq_list_def fsort_def Let_def
+ apply safe
+ subgoal by (simp add: rev_map[symmetric] RBT.distinct_entries)
+ subgoal by (auto simp: rev_map)
+ subgoal by (simp add: RBT.lookup_in_tree[symmetric])
+ subgoal by (simp add: RBT.lookup_in_tree[symmetric])
+ subgoal by (simp add: RBT.lookup_in_tree[symmetric])
+ done
+ (*>*)
+
+
+ section \<open>Code Generation\<close>
+ text \<open>Now we can use Isabelle/HOL's code generator to actually extract
+ functional code from our Isabelle formalization.
- lemma [code]: "fsort ws = fsort2 ws"
- unfolding fsort_def fsort2_def
- by simp
-
- definition fsort_string :: "String.literal list \<Rightarrow> String.literal list"
+ First, we derive a specialized versions with strings:
+ \<close>
+ definition fsort_string :: "String.literal list \<Rightarrow> (String.literal \<times> nat) list"
where "fsort_string \<equiv> fsort"
- export_code fsort_string in SML module_name Fsort file "export.sml"
+ text \<open>Then we can use the code generator in different ways.\<close>
+
+ text \<open>By the value command\<close>
+ value [code] "fsort_string [STR ''foo'', STR ''bar'', STR ''foo'', STR ''bara'']"
-
- value "fsort_string [STR ''foo'', STR ''bar'', STR ''foo'', STR ''bar'']"
-
- SML_file "export.sml"
- SML_file "fsort.sml"
-
+ text \<open>Export code to file\<close>
+ export_code fsort_string in SML module_name Fsort file "export.sml"
+ text \<open>We can load the file into JEdit's ML IDE: \<close>
+ SML_file "export.sml"
+ text \<open>And use it from within some wrapper code for parsing a corpus and printing the result: \<close>
+ SML_file "fsort.sml"
+
+ text \<open>Use code directly with Isabelle's builtin ML interpreter
+ \begin{verbatim}
+ ML_val {* see template file *}
+ \end{verbatim}
+ \<close>
+ (* Directly as Isabelle/ML command *)
+(*<*)
ML_val \<open>
(* Read file to string *)
fun file_to_string name = let
@@ -104,11 +248,16 @@
fun fs fname = @{code fsort_string}
(String.tokens (not o Char.isAlpha) (String.map (Char.toLower) (file_to_string fname)))
- val r = fs "/home/lammich/MASC-3.0.0/data/written/non-fiction/CUP2.txt"
+ val r1 = fs "/home/lammich/MASC-3.0.0/data/written/non-fiction/CUP2.txt"
+ val r2 = fs "/home/lammich/MASC-3.0.0/data/written/twitter/tweets2.txt"
\<close>
-
-
-
-
+(*>*)
+
+ text \<open>The code generator also supports other target languages\<close>
+ export_code fsort_string in Haskell
+ export_code fsort_string in Scala
+ export_code fsort_string in OCaml
+(*<*)
end
+(*>*)
--- a/Exercises/ex14/export.sml Sat Jul 29 19:34:17 2017 +0200
+++ b/Exercises/ex14/export.sml Sat Jul 29 19:34:29 2017 +0200
@@ -1,16 +1,10 @@
structure Fsort : sig
- val fsort_string : string list -> string list
+ type nat
+ val fsort_string : string list -> (string * nat) list
end = struct
datatype nat = Nat of IntInf.int;
-val zero_nata : nat = Nat (0 : IntInf.int);
-
-type 'a zero = {zero : 'a};
-val zero = #zero : 'a zero -> 'a;
-
-val zero_nat = {zero = zero_nata} : nat zero;
-
fun integer_of_nat (Nat x) = x;
fun less_eq_nat m n = IntInf.<= (integer_of_nat m, integer_of_nat n);
@@ -38,12 +32,6 @@
val linorder_nat = {order_linorder = order_nat} : nat linorder;
-type 'a equal = {equal : 'a -> 'a -> bool};
-val equal = #equal : 'a equal -> 'a -> 'a -> bool;
-
-val equal_literal = {equal = (fn a => fn b => ((a : string) = b))} :
- string equal;
-
val ord_literal =
{less_eq = (fn a => fn b => ((a : string) <= b)),
less = (fn a => fn b => ((a : string) < b))}
@@ -69,8 +57,6 @@
datatype ('b, 'a) rbt = RBT of ('b, 'a) rbta;
-fun eq A_ a b = equal A_ a b;
-
fun plus_nat m n = Nat (IntInf.+ (integer_of_nat m, integer_of_nat n));
val one_nat : nat = Nat (1 : IntInf.int);
@@ -85,8 +71,10 @@
fun equal_nat m n = (((integer_of_nat m) : IntInf.int) = (integer_of_nat n));
+val zero_nat : nat = Nat (0 : IntInf.int);
+
fun nth (x :: xs) n =
- (if equal_nat n zero_nata then x else nth xs (minus_nat n one_nat));
+ (if equal_nat n zero_nat then x else nth xs (minus_nat n one_nat));
fun fold f (x :: xs) s = fold f xs (f x s)
| fold f [] s = s;
@@ -255,9 +243,18 @@
rbt_lookup ((ord_preorder o preorder_order o order_linorder) A_)
(impl_of A_ x);
-fun incr2 A_ k m =
- (case lookup A_ m k of NONE => insert A_ k one_nat m
- | SOME v => insert A_ k (suc v) m);
+fun gen_entries kvts (Branch (c, l, k, v, r)) =
+ gen_entries (((k, v), r) :: kvts) l
+ | gen_entries ((kv, t) :: kvts) Empty = kv :: gen_entries kvts t
+ | gen_entries [] Empty = [];
+
+fun entriesa x = gen_entries [] x;
+
+fun entries A_ x = entriesa (impl_of A_ x);
+
+fun incr2 A_ k t =
+ (case lookup A_ t k of NONE => insert A_ k one_nat t
+ | SOME f => insert A_ k (plus_nat f one_nat) t);
fun freq2 A_ ws = fold (incr2 A_) ws (empty A_);
@@ -294,7 +291,7 @@
fun gen_length n (x :: xs) = gen_length (suc n) xs
| gen_length n [] = n;
-fun size_list x = gen_length zero_nata x;
+fun size_list x = gen_length zero_nat x;
fun part B_ f pivot (x :: xs) =
let
@@ -328,25 +325,10 @@
sort_key B_ f lts @ eqs @ sort_key B_ f gts
end);
-fun abs_freq1 B_ m k = (case m k of NONE => zero B_ | SOME v => v);
-
-fun member A_ [] y = false
- | member A_ (x :: xs) y = eq A_ x y orelse member A_ xs y;
-
-fun remdups A_ [] = []
- | remdups A_ (x :: xs) =
- (if member A_ xs x then remdups A_ xs else x :: remdups A_ xs);
+fun snd (x1, x2) = x2;
-fun fsort2 (A1_, A2_) ws =
- let
- val t = freq2 A2_ ws;
- in
- rev (sort_key linorder_nat (abs_freq1 zero_nat (lookup A2_ t))
- (remdups A1_ ws))
- end;
+fun fsort A_ ws = rev (sort_key linorder_nat snd (entries A_ (freq2 A_ ws)));
-fun fsort (A1_, A2_) ws = fsort2 (A1_, A2_) ws;
-
-fun fsort_string x = fsort (equal_literal, linorder_literal) x;
+fun fsort_string x = fsort linorder_literal x;
end; (*struct Fsort*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/ex14/tmpl14.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,160 @@
+text {* \ExerciseSheet{14}{28.~7.~2017} *}
+theory tmpl14
+imports
+ "~~/src/HOL/Library/Multiset"
+ "~~/src/HOL/Library/Code_Target_Nat"
+ "~~/src/HOL/Library/RBT"
+ "~~/src/HOL/Library/Char_ord"
+ "~~/src/HOL/Library/Code_Char"
+begin
+
+
+ text \<open>
+ \Exercise{Word Frequency --- Down to ML code}
+
+ Your task is to develop a program that reads a corpus and
+ prints the words in the corpus together with their frequencies,
+ sorted by descending frequencies.
+
+ Except input and output, your program shall be formalized in Isabelle/HOL.
+ A corpus is described as @{typ \<open>'a list\<close>},
+ and the result is described by @{typ \<open>('a \<times> nat) list\<close>}
+ \<close>
+
+ text \<open>The frequency of a word in a corpus can be specified as:\<close>
+ definition freq :: "'a list \<Rightarrow> 'a \<Rightarrow> nat"
+ where "freq ws = count (mset ws)"
+
+ text \<open>
+ Specify a predicate that characterizes a correct result.
+ Note: If words have the same frequency, any order is allowed.
+ \<close>
+ consts is_freq_list :: "'a list \<Rightarrow> ('a \<times> nat) list \<Rightarrow> bool"
+
+ text \<open>Tests:\<close>
+ lemma \<open>is_freq_list [''a'',''b'',''b'',''c'',''c''] [(''c'',2),(''b'',2),(''a'',1)]\<close>
+ oops
+ lemma \<open>is_freq_list [''a'',''b'',''b'',''c'',''c''] [(''b'',2),(''c'',2),(''a'',1)]\<close>
+ oops
+ lemma \<open>\<not>is_freq_list [''a'',''b'',''b'',''c'',''c''] [(''b'',2),(''c'',3),(''a'',1)]\<close>
+ oops
+ lemma \<open>\<not>is_freq_list [''a'',''b'',''b'',''c'',''c''] [(''a'',1),(''c'',2),(''b'',2)]\<close>
+ oops
+ lemma \<open>\<not>is_freq_list [''a'',''b'',''b'',''c'',''c''] [(''b'',2),(''c'',2),(''b'',2),(''a'',1)]\<close>
+ oops
+
+
+ section \<open>Refinement \#1\<close>
+ (* Refinement #1: Fold over word list and maintain current frequency *)
+ text \<open>Compute a function from words to their frequency by folding
+ over the corpus, starting with @{term \<open>\<lambda>_. 0::nat\<close>}. \<close>
+
+ consts incr1 :: "'a \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> 'a \<Rightarrow> nat"
+ definition "freq1 ws \<equiv> fold incr1 ws (\<lambda>_. 0)"
+
+ lemma freq1_correct[simp]: "freq1 ws = freq ws"
+ oops
+
+ section \<open>Refinement \#2\<close>
+ (* Refinement #2: Frequency map by RBT *)
+ text \<open>
+ Use red black trees to implement the mapping from words to frequencies.
+ Words that do not occur in the corpus must not be mapped!
+
+ Use the RBT implementation from \<open>HOL/Library/RBT\<close>!
+ It provides, e.g., @{const RBT.empty}, @{const RBT.lookup}, @{const RBT.insert}.
+ \<close>
+ definition abs_fm :: "('a::linorder,nat) rbt \<Rightarrow> 'a \<Rightarrow> nat" where
+ "abs_fm t k \<equiv> case RBT.lookup t k of None \<Rightarrow> 0 | Some v \<Rightarrow> v"
+ definition inv_fm :: "('a::linorder,nat) rbt \<Rightarrow> bool" where
+ "inv_fm t \<equiv> (0 \<notin> ran (RBT.lookup t))"
+
+ lemma empty2_correct[simp]:
+ "abs_fm RBT.empty = (\<lambda>_. 0)" "inv_fm RBT.empty"
+ oops
+
+ definition incr2 :: "'a::linorder \<Rightarrow> ('a, nat) rbt \<Rightarrow> ('a, nat) rbt"
+ where "incr2 k t = t"
+
+ lemma incr2_correct[simp]:
+ "inv_fm t \<Longrightarrow> abs_fm (incr2 k t) = incr1 k (abs_fm t)"
+ "inv_fm t \<Longrightarrow> inv_fm (incr2 k t)"
+ oops
+
+ text \<open> Now we have refined the operations, we can refine the algorithm that uses the operations:\<close>
+ definition "freq2 ws \<equiv> fold incr2 ws RBT.empty"
+
+ lemma freq2_correct[simp]: "abs_fm (freq2 ws) = freq1 ws" "inv_fm (freq2 ws)"
+ oops
+
+
+ subsection \<open>Extracting Result from RBT\<close>
+ text \<open>Extract the desired result
+ --- list of pairs of words and their frequencies, sorted by frequency ---
+ from the red black tree. Use @{const RBT.entries}.
+ \<close>
+
+ definition fsort :: "'a::linorder list \<Rightarrow> ('a \<times> nat) list"
+ where "fsort ws \<equiv> []"
+
+
+ text \<open>Prove that your function is correct.
+ Hint: You will need to prove some auxiliary lemmas on standard list functions.
+ Use \<open>find_theorems\<close> to search for existing lemmas.
+ Hint: A lemma of the form @{text \<open>RBT.lookup (freq2 ws) w = Some f \<longleftrightarrow> \<dots>\<close>},
+ derived from @{text \<open>freq2_correct freq1_correct\<close>} may be useful!
+ \<close>
+ lemma fsort_correct: "is_freq_list ws (fsort ws)"
+ oops
+
+
+ section \<open>Code Generation\<close>
+ text \<open>Now we can use Isabelle/HOL's code generator to actually extract
+ functional code from our Isabelle formalization.
+
+ First, we derive a specialized versions with strings:
+ \<close>
+ definition fsort_string :: "String.literal list \<Rightarrow> (String.literal \<times> nat) list"
+ where "fsort_string \<equiv> fsort"
+
+ text \<open>Then we can use the code generator in different ways.\<close>
+
+ text \<open>By the value command\<close>
+ value [code] "fsort_string [STR ''foo'', STR ''bar'', STR ''foo'', STR ''bara'']"
+
+ text \<open>Export code to file\<close>
+ export_code fsort_string in SML module_name Fsort file "export.sml"
+ text \<open>We can load the file into JEdit's ML IDE: \<close>
+ SML_file "export.sml"
+ text \<open>And use it from within some wrapper code for parsing a corpus and printing the result: \<close>
+ SML_file "fsort.sml"
+
+ text \<open>Use code directly with Isabelle's builtin ML interpreter
+ \begin{verbatim}
+ ML_val {* see template file *}
+ \end{verbatim}
+ \<close>
+ (* Directly as Isabelle/ML command *)
+ ML_val \<open>
+ (* Read file to string *)
+ fun file_to_string name = let
+ val f = TextIO.openIn name
+ val s = TextIO.inputAll f
+ val _ = TextIO.closeIn f
+ in s end
+
+ fun fs fname = @{code fsort_string}
+ (String.tokens (not o Char.isAlpha) (String.map (Char.toLower) (file_to_string fname)))
+
+ val r1 = fs "/home/lammich/MASC-3.0.0/data/written/non-fiction/CUP2.txt"
+ val r2 = fs "/home/lammich/MASC-3.0.0/data/written/twitter/tweets2.txt"
+ \<close>
+
+ text \<open>The code generator also supports other target languages\<close>
+ export_code fsort_string in Haskell
+ export_code fsort_string in Scala
+ export_code fsort_string in OCaml
+
+(*<*)
+end
+(*>*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/ex14/tut14.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,211 @@
+text {* \ExerciseSheet{14}{28.~7.~2017} *}
+theory tut14
+imports
+ "~~/src/HOL/Library/Multiset"
+ "~~/src/HOL/Library/Code_Target_Nat"
+ "~~/src/HOL/Library/RBT"
+ "~~/src/HOL/Library/Char_ord"
+ "~~/src/HOL/Library/Code_Char"
+begin
+
+
+ text \<open>
+ \Exercise{Word Frequency --- Down to ML code}
+
+ Your task is to develop a program that reads a corpus and
+ prints the words in the corpus together with their frequencies,
+ sorted by descending frequencies.
+
+ Except input and output, your program shall be formalized in Isabelle/HOL.
+ A corpus is described as @{typ \<open>'a list\<close>},
+ and the result is described by @{typ \<open>('a \<times> nat) list\<close>}
+ \<close>
+
+ text \<open>The frequency of a word in a corpus can be specified as:\<close>
+ definition freq :: "'a list \<Rightarrow> 'a \<Rightarrow> nat"
+ where "freq ws = count (mset ws)"
+
+ text \<open>
+ Specify a predicate that characterizes a correct result.
+ Note: If words have the same frequency, any order is allowed.
+ \<close>
+ definition is_freq_list :: "'a list \<Rightarrow> ('a \<times> nat) list \<Rightarrow> bool"
+ where
+ "is_freq_list ws fl \<equiv>
+ distinct (map fst fl)
+ \<and> (\<forall>w f. (w,f)\<in>set fl \<longleftrightarrow> freq ws w = f \<and> f>0 )
+ \<and> sorted (rev (map snd fl))
+ "
+
+ text \<open>Tests:\<close>
+ lemma \<open>is_freq_list [''a'',''b'',''b'',''c'',''c''] [(''c'',2),(''b'',2),(''a'',1)]\<close>
+ by (auto simp: is_freq_list_def freq_def)
+ lemma \<open>is_freq_list [''a'',''b'',''b'',''c'',''c''] [(''b'',2),(''c'',2),(''a'',1)]\<close>
+ by (auto simp: is_freq_list_def freq_def)
+ lemma \<open>\<not>is_freq_list [''a'',''b'',''b'',''c'',''c''] [(''b'',2),(''c'',3),(''a'',1)]\<close>
+ by (auto simp: is_freq_list_def freq_def)
+ lemma \<open>\<not>is_freq_list [''a'',''b'',''b'',''c'',''c''] [(''a'',1),(''c'',2),(''b'',2)]\<close>
+ by (auto simp: is_freq_list_def freq_def)
+ lemma \<open>\<not>is_freq_list [''a'',''b'',''b'',''c'',''c''] [(''b'',2),(''c'',2),(''b'',2),(''a'',1)]\<close>
+ by (auto simp: is_freq_list_def freq_def)
+
+
+ section \<open>Refinement \#1\<close>
+ (* Refinement #1: Fold over word list and maintain current frequency *)
+ text \<open>Compute a function from words to their frequency by folding
+ over the corpus, starting with @{term \<open>\<lambda>_. 0::nat\<close>}. \<close>
+
+ term "Map.empty"
+ term "f(x:=y)"
+
+ definition incr1 :: "'a \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> 'a \<Rightarrow> nat"
+ where "incr1 w fm \<equiv> fm(w:=fm w + 1)"
+ definition "freq1 ws \<equiv> fold incr1 ws (\<lambda>_. 0)"
+
+ lemma freq1_correct[simp]: "freq1 ws = freq ws"
+ proof -
+ have "fold incr1 ws fm = (\<lambda>w. count (mset ws) w + fm w)" for fm
+ by (induction ws arbitrary: fm) (auto simp: incr1_def)
+ from this[of "\<lambda>_. 0"] show ?thesis
+ unfolding freq1_def freq_def by auto
+ qed
+
+
+ section \<open>Refinement \#2\<close>
+ (* Refinement #2: Frequency map by RBT *)
+ text \<open>
+ Use red black trees to implement the mapping from words to frequencies.
+ Words that do not occur in the corpus must not be mapped!
+
+ Use the RBT implementation from \<open>HOL/Library/RBT\<close>!
+ It provides, e.g., @{const RBT.empty}, @{const RBT.lookup}, @{const RBT.insert}.
+ \<close>
+ definition abs_fm :: "('a::linorder,nat) rbt \<Rightarrow> 'a \<Rightarrow> nat" where
+ "abs_fm t k \<equiv> case RBT.lookup t k of None \<Rightarrow> 0 | Some v \<Rightarrow> v"
+ definition inv_fm :: "('a::linorder,nat) rbt \<Rightarrow> bool" where
+ "inv_fm t \<equiv> (0 \<notin> ran (RBT.lookup t))"
+
+ lemma empty2_correct[simp]:
+ "abs_fm RBT.empty = (\<lambda>_. 0)" "inv_fm RBT.empty"
+ unfolding abs_fm_def inv_fm_def by auto
+
+ term RBT.lookup term RBT.insert
+
+ definition incr2 :: "'a::linorder \<Rightarrow> ('a, nat) rbt \<Rightarrow> ('a, nat) rbt"
+ where "incr2 k t = (
+ case RBT.lookup t k of
+ None \<Rightarrow> RBT.insert k 1 t
+ | Some f \<Rightarrow> RBT.insert k (f+1) t)"
+
+ lemma incr2_correct[simp]:
+ "inv_fm t \<Longrightarrow> abs_fm (incr2 k t) = incr1 k (abs_fm t)"
+ "inv_fm t \<Longrightarrow> inv_fm (incr2 k t)"
+ unfolding inv_fm_def abs_fm_def incr2_def incr1_def
+ by (auto simp: ran_def split: option.split)
+
+ text \<open> Now we have refined the operations, we can refine the algorithm that uses the operations:\<close>
+ definition "freq2 ws \<equiv> fold incr2 ws RBT.empty"
+
+ lemma freq2_correct[simp]:
+ shows
+ "abs_fm (freq2 ws) = freq1 ws" (is ?G1)
+ and "inv_fm (freq2 ws)" (is ?G2)
+ proof -
+ have "abs_fm (fold incr2 ws t) = fold incr1 ws (abs_fm t)
+ \<and> inv_fm (fold incr2 ws t)" if "inv_fm t" for t
+ using that
+ by (induction ws arbitrary: t) (auto)
+ from this show ?G1 ?G2
+ unfolding freq2_def freq1_def
+ by auto
+ qed
+
+ subsection \<open>Extracting Result from RBT\<close>
+ text \<open>Extract the desired result
+ --- list of pairs of words and their frequencies, sorted by frequency ---
+ from the red black tree. Use @{const RBT.entries}.
+ \<close>
+
+ term "RBT.entries (freq2 ws)"
+ term sort_key
+ definition fsort :: "'a::linorder list \<Rightarrow> ('a \<times> nat) list"
+ where "fsort ws \<equiv> rev (sort_key snd (RBT.entries (freq2 ws)))"
+
+
+ text \<open>Prove that your function is correct.
+ Hint: You will need to prove some auxiliary lemmas on standard list functions.
+ Use \<open>find_theorems\<close> to search for existing lemmas.
+ Hint: A lemma of the form @{text \<open>RBT.lookup (freq2 ws) w = Some f \<longleftrightarrow> \<dots>\<close>},
+ derived from @{text \<open>freq2_correct freq1_correct\<close>} may be useful!
+ \<close>
+
+ lemma [simp]: "(w, f) \<in> set (RBT.entries (freq2 ws)) \<longleftrightarrow> freq ws w = f \<and> f>0"
+ using freq2_correct(1)[of ws, THEN fun_cong[where x=w]]
+ using freq2_correct(2)[of ws]
+ unfolding abs_fm_def inv_fm_def
+ apply (auto split: option.splits simp: ran_def lookup_in_tree[symmetric])
+ using neq0_conv by fastforce
+
+
+ lemma fsort_correct: "is_freq_list ws (fsort ws)"
+ unfolding is_freq_list_def
+ apply (intro conjI)
+ subgoal unfolding fsort_def
+ apply (simp add: rev_map[symmetric])
+ by (metis RBT.distinct_entries distinct_map distinct_sort set_sort)
+ subgoal unfolding fsort_def
+ by auto
+ subgoal unfolding fsort_def
+ by (auto simp: rev_map)
+ done
+
+ section \<open>Code Generation\<close>
+ text \<open>Now we can use Isabelle/HOL's code generator to actually extract
+ functional code from our Isabelle formalization.
+
+ First, we derive a specialized versions with strings:
+ \<close>
+ definition fsort_string :: "String.literal list \<Rightarrow> (String.literal \<times> nat) list"
+ where "fsort_string \<equiv> fsort"
+
+ text \<open>Then we can use the code generator in different ways.\<close>
+
+ text \<open>By the value command\<close>
+ value [code] "fsort_string [STR ''foo'', STR ''bar'', STR ''foo'', STR ''bara'']"
+
+ text \<open>Export code to file\<close>
+ export_code fsort_string in SML module_name Fsort file "export.sml"
+ text \<open>We can load the file into JEdit's ML IDE: \<close>
+ SML_file "export.sml"
+ text \<open>And use it from within some wrapper code for parsing a corpus and printing the result: \<close>
+ SML_file "fsort.sml"
+
+ text \<open>Use code directly with Isabelle's builtin ML interpreter
+ \begin{verbatim}
+ ML_val {* see template file *}
+ \end{verbatim}
+ \<close>
+ (* Directly as Isabelle/ML command *)
+ ML_val \<open>
+ (* Read file to string *)
+ fun file_to_string name = let
+ val f = TextIO.openIn name
+ val s = TextIO.inputAll f
+ val _ = TextIO.closeIn f
+ in s end
+
+ fun fs fname = @{code fsort_string}
+ (String.tokens (not o Char.isAlpha) (String.map (Char.toLower) (file_to_string fname)))
+
+ val r1 = fs "/home/lammich/MASC-3.0.0/data/written/non-fiction/CUP2.txt"
+ val r2 = fs "/home/lammich/MASC-3.0.0/data/written/twitter/tweets2.txt"
+ \<close>
+
+ text \<open>The code generator also supports other target languages\<close>
+ export_code fsort_string in Haskell
+ export_code fsort_string in Scala
+ export_code fsort_string in OCaml
+
+(*<*)
+end
+(*>*)
Binary file Exercises/exam.pdf has changed
--- a/Exercises/exam/Q_Balanced_Insert.thy Sat Jul 29 19:34:17 2017 +0200
+++ b/Exercises/exam/Q_Balanced_Insert.thy Sat Jul 29 19:34:29 2017 +0200
@@ -27,7 +27,7 @@
list into an initially empty search tree:
\<close>
definition from_list :: "'a::linorder list \<Rightarrow> 'a tree" where
- "from_list l = fold ins l \<langle>\<rangle>"
+ "from_list l = fold ins l Leaf"
text \<open>
Your task is to specify a function @{term_type preprocess},
--- a/Exercises/exam/Q_Induction.thy Sat Jul 29 19:34:17 2017 +0200
+++ b/Exercises/exam/Q_Induction.thy Sat Jul 29 19:34:29 2017 +0200
@@ -21,7 +21,7 @@
"lvs (Nd r s t) = lvs r + lvs s + lvs t"
text \<open>There is a linear relationship between the size (\<open>sz\<close>) and the number of
-leaves (\<open>lvs\<close>) of a tree. Find this relationship and prove \<open>lfs t = a * sz t + b\<close>
+leaves (\<open>lvs\<close>) of a tree. Find this relationship and prove \<open>lvs t = a * sz t + b\<close>
for the correct \<open>a\<close> and \<open>b\<close>.\<close>
(*<*)
--- a/Exercises/exam/Q_Tree_Same_Struct.thy Sat Jul 29 19:34:17 2017 +0200
+++ b/Exercises/exam/Q_Tree_Same_Struct.thy Sat Jul 29 19:34:29 2017 +0200
@@ -83,3 +83,5 @@
end
(*>*)
text_raw \<open>\blankpage\<close>
+text_raw \<open>\blankpage\<close>
+text_raw \<open>\blankpage\<close>
Binary file Exercises/hwsol.pdf has changed
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Biendarra_Julian_julian.biendarra@tum.de_588/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+julian.biendarra@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Biendarra_Julian_julian.biendarra@tum.de_588/hw12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,67 @@
+(** Score: 5/5
+*)
+(*<*)
+theory hw12
+imports Complex_Main
+begin
+(*>*)
+
+locale Dyn_Tab
+begin
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) = (n \<le> l \<and> l \<le> 2 * n)"
+
+definition init :: tab where
+"init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+apply(cases t)
+apply(auto)
+done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = 2*(real n) - (real l)"
+
+text \<open>Careful: \<open>n\<close> and \<open>l\<close> are natural numbers.
+Thus \<open>2*n\<close> can be less than \<open>l\<close> (in which case \<open>2*n - l = 0\<close>!)
+because the invariant does not exclude this case, although it cannot arise.
+If you go through the proof of lemma \<open>a_ins\<close> in detail
+you will understand why this case also works out fine.\<close>
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+apply(auto)
+done
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+apply(cases t)
+apply(auto split: if_splits)
+done
+
+end
+
+(*<*)
+end
+(*>*)
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Biendarra_Julian_julian.biendarra@tum.de_588/testoutput.html Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,24 @@
+
+<html>
+<head>
+<link rel="stylesheet" href="https://maxcdn.bootstrapcdn.com/bootstrap/3.3.7/css/bootstrap.min.css">
+</head>
+<body>
+<h3>1/1 test cases passed!</h3>
+
+<table class="table">
+<thead>
+<tr>
+<th>Test Case</th><th>Passed</th>
+</tr>
+</thead>
+<tbody>
+
+<tr><td>check_build</td><td>Passed</td></tr>
+
+</tbody>
+</table>
+
+</body>
+</html>
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Erhard_Julian_ga48kog@mytum.de_604/Amortized_Examples.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,329 @@
+(** Score: 4/5
+ File contains Isabelle error, l.276: -1
+
+ don't submit files with errors!
+ why not just delete the irrelevant parts, such that I can find your
+ solution more quickly?
+
+*)
+section "Amortized Complexity: Classical Examples"
+
+theory Amortized_Examples
+imports Complex_Main
+begin
+
+subsection "Binary Counter"
+
+locale Counter
+begin
+
+type_synonym counter = "bool list"
+
+definition init :: counter where
+"init = []"
+
+fun incr :: "counter \<Rightarrow> counter" where
+"incr [] = [True]" |
+"incr (False # bs) = True # bs" |
+"incr (True # bs) = False # incr bs"
+
+fun t_incr :: "counter \<Rightarrow> nat" where
+"t_incr [] = 1" |
+"t_incr (False # bs) = 1" |
+"t_incr (True # bs) = t_incr bs + 1"
+
+definition \<Phi> :: "counter \<Rightarrow> int" where
+"\<Phi> bs = length(filter id bs)"
+
+lemma \<Phi>_non_neg: "\<Phi> bs \<ge> 0"
+by(simp add: \<Phi>_def)
+
+lemma \<Phi>_init: "\<Phi> init = 0"
+by(simp add: \<Phi>_def init_def)
+
+lemma a_incr: "t_incr bs + \<Phi>(incr bs) - \<Phi> bs = 2"
+apply(induction bs rule: incr.induct)
+apply (simp_all add: \<Phi>_def)
+done
+
+(* Application: Doing n increments *)
+
+fun n_increments where
+ "n_increments 0 bs = bs"
+| "n_increments (Suc n) bs = n_increments n (incr bs)"
+
+fun t_n_increments :: "nat \<Rightarrow> counter \<Rightarrow> nat" where
+ "t_n_increments 0 bs = 1"
+| "t_n_increments (Suc n) bs = t_incr bs + t_n_increments n (incr bs)"
+
+(* Generalized lemma with arbitrary start state.
+ Telescoping sum is done implicitly in this proof. *)
+lemma t_n_increments_aux:
+ "t_n_increments n bs = 1 + 2*n + \<Phi> bs - \<Phi> (n_increments n bs)"
+proof (induction n bs rule: n_increments.induct)
+ case (1 bs)
+ then show ?case by simp
+next
+ case (2 n bs)
+ then show ?case
+ using a_incr[of bs] by auto
+qed
+
+(* Estimation for initial state *)
+lemma t_n_increments: "t_n_increments n init \<le> 1+2*n"
+ using
+ t_n_increments_aux[of n init]
+ \<Phi>_non_neg[of "n_increments n init"] (* Exploit that potential cannot be negative *)
+ \<Phi>_init
+ by auto
+
+
+end
+
+
+subsection "Dynamic tables: insert only"
+
+locale Dyn_Tab
+begin
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) = (n \<le> l \<and> l \<le> 2 * n)"
+
+definition init :: tab where
+"init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+apply(cases t)
+apply(auto)
+done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = 2*(real n) - (real l)"
+
+text \<open>Careful: \<open>n\<close> and \<open>l\<close> are natural numbers.
+Thus \<open>2*n\<close> can be less than \<open>l\<close> (in which case \<open>2*n - l = 0\<close>!)
+because the invariant does not exclude this case, although it cannot arise.
+If you go through the proof of lemma \<open>a_ins\<close> in detail
+you will understand why this case also works out fine.\<close>
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+apply(auto)
+done
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+apply(cases t)
+apply(auto split: if_splits)
+done
+
+end
+
+
+subsection "Queue"
+
+locale Queue
+begin
+
+text \<open>Very simplified model, dequeue returns no result.\<close>
+(* REMARK: Using \<open>unit list\<close> would simplify model even more. *)
+
+type_synonym 'a queue = "'a list * 'a list"
+
+definition init :: "'a queue" where
+"init = ([],[])"
+
+fun enq :: "'a \<Rightarrow> 'a queue \<Rightarrow> 'a queue" where
+"enq x (xs,ys) = (x#xs, ys)"
+
+fun deq :: "'a queue \<Rightarrow> 'a queue" where
+"deq (xs,ys) = (if ys = [] then ([], tl(rev xs)) else (xs, tl ys))"
+
+text \<open>Time: we count only the number of newly allocated list cells.\<close>
+
+fun t_enq :: "'a \<Rightarrow> 'a queue \<Rightarrow> nat" where
+"t_enq x (xs,ys) = 1"
+
+text \<open>We assume that function \<open>rev\<close> has linear complexity.\<close>
+
+fun t_deq :: "'a queue \<Rightarrow> nat" where
+"t_deq (xs,ys) = (if ys = [] then length xs else 0)"
+
+fun \<Phi> :: "'a queue \<Rightarrow> int" where
+"\<Phi> (xs,ys) = length xs"
+
+lemma \<Phi>_non_neg: "\<Phi> q \<ge> 0"
+by(cases q) simp
+
+lemma \<Phi>_init: "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_enq: "t_enq x q + \<Phi>(enq x q) - \<Phi> q = 2"
+by(cases q) auto
+
+lemma a_deq: "t_deq q + \<Phi>(deq q) - \<Phi> q = 0"
+by(cases q) auto
+
+text \<open>
+ Estimating the cost of a sequence of enqueue and dequeue operations.
+ Note: The cost of a dequeue operation on the empty list is just 0, thus
+ that we do not care about the queue being actually non-empty if we
+ call \<open>deq\<close>.
+\<close>
+
+lemma "t_deq init = 0" by (auto simp: init_def)
+
+datatype 'a opr = ENQ 'a | DEQ
+
+fun execute :: "'a queue \<Rightarrow> 'a opr list \<Rightarrow> 'a queue" where
+ "execute q [] = q"
+| "execute q (ENQ x#ops) = execute (enq x q) ops"
+| "execute q (DEQ#ops) = execute (deq q) ops"
+
+fun t_execute :: "'a queue \<Rightarrow> 'a opr list \<Rightarrow> nat" where
+ "t_execute q [] = 0"
+| "t_execute q (ENQ x#ops) = t_enq x q + t_execute (enq x q) ops"
+| "t_execute q (DEQ#ops) = t_deq q + t_execute (deq q) ops"
+
+lemma "t_execute q ops \<le> 2*length ops + \<Phi> q - \<Phi> (execute q ops)"
+ apply (induction q ops rule: execute.induct)
+ apply auto
+ done
+
+(* Note: The simplifier re-proves a_enq and a_deq here, as the definitions of
+ the operations and potential are in the simpset by default.
+
+ This may be nice for simple properties, but for more complex properties,
+ it will just result in very deep unfolding, exploding the subgoals and thus
+ rendering them intractable.
+
+ Thus, we clean up the simpset to not unfold below our abstraction boundary,
+ which is the operations and the potential:
+*)
+lemmas [simp del] = \<Phi>.simps enq.simps deq.simps t_enq.simps t_deq.simps
+
+(* This proof structure will also work for more complicated estimates! *)
+lemma t_execute_aux: "t_execute q ops \<le> 2*length ops + \<Phi> q - \<Phi> (execute q ops)"
+proof (induction q ops rule: execute.induct)
+ case (1 q)
+ then show ?case by simp
+next
+ case (2 q x ops)
+ then show ?case using a_enq[of x q] by simp
+next
+ case (3 q ops)
+ then show ?case using a_deq[of q] by simp
+qed
+
+
+lemma t_execute: "t_execute init ops \<le> 2*length ops"
+ using t_execute_aux[of init ops] \<Phi>_non_neg[of "execute init ops"]
+ by (auto simp: \<Phi>_init)
+
+end
+
+
+subsection "Dynamic tables: insert and delete"
+
+locale Dyn_Tab2 = Dyn_Tab
+begin
+
+fun del :: "tab \<Rightarrow> tab" where
+"del (n,l) = (n-1, if n=1 then 0 else if 4*(n-1) < l then l div 2 else l)"
+
+fun t_del :: "tab \<Rightarrow> nat" where
+"t_del (n,l) = (if n=1 then 1 else if 4*(n-1) < l then n else 1)"
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = (if 2*n < l then l/2 - n else 2*n - l)"
+
+lemma \<Phi>_non_negative: "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+by(cases t) auto
+
+lemma \<Phi>_init: "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+by (cases t) (auto split: if_splits)
+
+lemma a_del: "invar t \<Longrightarrow> t_del t + \<Phi>(del t) - \<Phi> t \<le> 2"
+by (cases t) (auto split: if_splits)
+
+lemma init_invar: "invar init" by (auto simp: init_def)
+lemma ins_invar: "invar t \<Longrightarrow> invar (ins t)"
+ by (cases t) (auto)
+
+lemma del_invar: "invar t \<Longrightarrow> invar (del t)"
+ by (cases t) (auto)
+
+text \<open>Establish reasonable abstraction boundary\<close>
+lemmas [simp del] = ins.simps t_ins.simps del.simps t_del.simps \<Phi>.simps
+
+datatype opr = INS | DEL
+
+fun execute :: "tab \<Rightarrow> opr list \<Rightarrow> tab" where
+ "execute t [] = t"
+| "execute t (INS#ops) = execute (ins t) ops"
+| "execute t (DEL#ops) = execute (del t) ops"
+
+fun t_execute :: "tab \<Rightarrow> opr list \<Rightarrow> nat" where
+ "t_execute t [] = 1"
+| "t_execute t (INS#ops) = t_ins t + t_execute (ins t) ops"
+| "t_execute t (DEL#ops) = t_del t + t_execute (del t) ops"
+
+lemma execute_invar: "invar t \<Longrightarrow> invar (execute t ops)"
+ by (induction t ops rule: execute.induct) (auto simp: ins_invar del_invar)
+
+(*
+ To use the amortized complexity estimate here,
+ we have to additionally keep track of the invariant.
+*)
+lemma t_execute_aux: "invar t \<Longrightarrow> t_execute t ops \<le> 1 + 3*length ops + \<Phi> t - \<Phi> (execute t ops)"
+proof (induction t ops rule: execute.induct)
+ case (1 t)
+ then show ?case by simp
+next
+ case (2 t ops)
+ then show ?case
+ using a_ins[of t]
+ by (auto simp: ins_invar)
+
+next
+ case (3 t ops)
+ then show ?case
+ using a_del[of t]
+ by (auto simp: del_invar)
+
+qed
+
+lemma t_execute: "t_execute init ops \<le> 1 + 3*length ops"
+ using
+ t_execute_aux[of init ops]
+ \<Phi>_non_negative[OF execute_invar, of init ops]
+ init_invar
+ \<Phi>_init
+ by auto
+
+end
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Erhard_Julian_ga48kog@mytum.de_604/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+ga48kog@mytum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Erhard_Julian_ga48kog@mytum.de_604/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.LIY5H1Dhik'
+Files in /tmp/eval-604-OvvpRv: Amortized_Examples.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Groer_Markus_markus.grosser@tum.de_597/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+markus.grosser@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Groer_Markus_markus.grosser@tum.de_597/hw12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,65 @@
+(** Score: 5/5
+*)
+theory hw12
+ imports Complex_Main
+begin
+
+locale Dyn_Tab
+begin
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+ "invar (n,l) = ((n \<ge> l/2) \<and> (n \<le> l))" \<comment> \<open>
+ I spent more time thinking about whether I missed something than on the actual homework @{text \<dots>}
+\<close>
+
+definition init :: tab where
+ "init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+ "ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+ "t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+ by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+ apply(cases t)
+ apply(auto)
+ done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+ "\<Phi> (n,l) = 2*(real n) - real l"
+
+text \<open>Careful: \<open>n\<close> and \<open>l\<close> are natural numbers.
+Thus \<open>2*n\<close> can be less than \<open>l\<close> (in which case \<open>2*n - l = 0\<close>!)
+because the invariant does not exclude this case, although it cannot arise.
+If you go through the proof of lemma \<open>a_ins\<close> in detail
+you will understand why this case also works out fine.\<close>
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+ apply(cases t)
+ apply(auto)
+ done
+
+lemma "\<Phi> init = 0"
+ by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+ apply(cases t)
+ apply(auto split: if_splits)
+ done
+
+end
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Groer_Markus_markus.grosser@tum.de_597/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.4lbWLNqwAj'
+Files in /tmp/eval-597-Ar4dsM: hw12.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Haslbeck_Maximilian_max.haslbeck@mytum.de_582/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+max.haslbeck@mytum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Haslbeck_Maximilian_max.haslbeck@mytum.de_582/ex12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,53 @@
+(** Score: 5/5
+*)
+theory ex12
+imports "../Thys/Amortized_Examples"
+begin
+type_synonym tab = "nat \<times> nat"
+
+locale Dyn_Tab
+begin
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) = (n \<le> l \<and> 2 * n \<ge> l)"
+
+definition init :: tab where
+"init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+apply(cases t)
+apply(auto)
+done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+ "\<Phi> (n,l) = 2 * (real n) - real l"
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+apply(auto)
+done
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+apply(cases t)
+apply(auto split: if_splits)
+done
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Haslbeck_Maximilian_max.haslbeck@mytum.de_582/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.zCkuBpf1yB'
+Files in /tmp/eval-582-7BZw2v: ex12.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Hellauer_Fabian_f.hellauer@tum.de_610/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+f.hellauer@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Hellauer_Fabian_f.hellauer@tum.de_610/hw12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,63 @@
+(** Score: 5/5
+*)
+theory hw12
+imports Complex_Main
+begin
+
+locale Dyn_Tab
+begin
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) = (l \<in> {n..2*n})"
+
+definition init :: tab where
+"init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+apply(cases t)
+apply(auto)
+done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = 2*(real n) - real l"
+
+text \<open>Careful: \<open>n\<close> and \<open>l\<close> are natural numbers.
+Thus \<open>2*n\<close> can be less than \<open>l\<close> (in which case \<open>2*n - l = 0\<close>!)
+because the invariant does not exclude this case, although it cannot arise.
+If you go through the proof of lemma \<open>a_ins\<close> in detail
+you will understand why this case also works out fine.\<close>
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+apply(auto)
+done
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+apply(cases t)
+apply(auto split: if_splits)
+done
+
+end
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Hellauer_Fabian_f.hellauer@tum.de_610/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.WpfyOmWl7j'
+Files in /tmp/eval-610-nYh1sB: hw12.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Hu_Shuwei_shuwei.hu@tum.de_573/Homework12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,47 @@
+(** Score: 5/5
+*)
+(* Shuwei Hu *)
+
+theory Homework12
+ imports Complex_Main
+begin
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) \<longleftrightarrow> (n \<le> l) \<and> (l \<le> n+n)"
+
+definition init :: tab where
+"init = (0,0)"
+
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+apply(cases t)
+apply(auto)
+done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = 2*(real n) - (real l)"
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+apply(auto)
+done
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+apply(cases t)
+apply(auto split: if_splits)
+done
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Hu_Shuwei_shuwei.hu@tum.de_573/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+shuwei.hu@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Hu_Shuwei_shuwei.hu@tum.de_573/testoutput.html Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,24 @@
+
+<html>
+<head>
+<link rel="stylesheet" href="https://maxcdn.bootstrapcdn.com/bootstrap/3.3.7/css/bootstrap.min.css">
+</head>
+<body>
+<h3>1/1 test cases passed!</h3>
+
+<table class="table">
+<thead>
+<tr>
+<th>Test Case</th><th>Passed</th>
+</tr>
+</thead>
+<tbody>
+
+<tr><td>check_build</td><td>Passed</td></tr>
+
+</tbody>
+</table>
+
+</body>
+</html>
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Hutzler_Matthias_ga95luy@mytum.de_602/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+ga95luy@mytum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Hutzler_Matthias_ga95luy@mytum.de_602/ex12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,87 @@
+(** Score: 5/5
+ insert was sufficient for this homework.
+*)
+theory ex12
+ imports Complex_Main
+begin
+
+
+ (* Homework 12 *)
+
+(* locale Dyn_Tab copied from Thys/Amortized_Examples and then modified *)
+
+locale Dyn_Tab
+begin
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) \<longleftrightarrow> n \<le> l \<and> l \<le> 2*n"
+
+definition init :: tab where
+"init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+apply(cases t)
+apply(auto)
+done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = 2*(real n) - (real l)"
+
+text \<open>Careful: \<open>n\<close> and \<open>l\<close> are natural numbers.
+Thus \<open>2*n\<close> can be less than \<open>l\<close> (in which case \<open>2*n - l = 0\<close>!)
+because the invariant does not exclude this case, although it cannot arise.
+If you go through the proof of lemma \<open>a_ins\<close> in detail
+you will understand why this case also works out fine.\<close>
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+apply(auto)
+done
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+apply(cases t)
+apply(auto split: if_splits)
+done
+
+end
+
+
+ (* I am not sure about the scope of this homework.
+ Should the "locale Dyn_Tab2" part be adjusted too? At first I thought so,
+ but then I realised that the definition for the potential \<Phi> from the
+ exercise sheet makes negative potentials unavoidable in presence of the
+ del function from Dyn_Tab2. *)
+
+locale Dyn_Tab2 = Dyn_Tab
+begin
+
+fun del :: "tab \<Rightarrow> tab" where
+"del (n,l) = (n-1, if n=1 then 0 else if 4*(n-1) < l then l div 2 else l)"
+
+(* The definitions of ins and del lead to negative potentials
+ (\<Phi> as given on the exercise sheet) in certain cases. *)
+lemma negative_\<Phi>_example: "\<Phi> (del (del (ins (ins (ins init))))) = -2"
+ unfolding init_def
+ by simp
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Hutzler_Matthias_ga95luy@mytum.de_602/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.vU51WwZNWG'
+Files in /tmp/eval-602-FLVy1u: ex12.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Keinholz_Jonas_jonas.keinholz@tum.de_572/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+jonas.keinholz@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Keinholz_Jonas_jonas.keinholz@tum.de_572/hw12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,47 @@
+(** Score: 5/5
+*)
+theory hw12
+ imports Complex_Main
+begin
+
+section \<open>Homework 12 Dynamic Tables with real-valued Potential\<close>
+
+locale Dyn_Tab
+begin
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+ "invar (n,l) \<longleftrightarrow> n \<le> l \<and> l \<le> 2 * n"
+
+definition init :: tab where
+ "init = (0,0)"
+
+fun ins :: "tab \<Rightarrow> tab" where
+ "ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+fun t_ins :: "tab \<Rightarrow> nat" where
+ "t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+ by (simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar (ins t)"
+ by (cases t) auto
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+ "\<Phi> (n,l) = 2 * (real n) - (real l)"
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+ by (cases t) auto
+
+lemma "\<Phi> init = 0"
+ unfolding init_def by auto
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+ by (cases t) (auto split: if_splits)
+
+end
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Keinholz_Jonas_jonas.keinholz@tum.de_572/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.iRwpDnvnBK'
+Files in /tmp/eval-572-JTIOJv: hw12.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Kurz_Friedrich_friedrich.kurz@tum.de_606/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+friedrich.kurz@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Kurz_Friedrich_friedrich.kurz@tum.de_606/hw_12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,62 @@
+(** Score: 5/5
+*)
+theory hw_12
+ imports "../../../Public/Thys/Amortized_Examples"
+begin
+
+locale Dyn_Tab
+begin
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) = (real n \<le> real l & 2*(real n) \<ge> real l)"
+
+definition init :: tab where
+"init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+ apply(cases t)
+ apply(auto)
+ done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+ "\<Phi> (n, l) = 2 * (real n) - real l"
+
+text \<open>Careful: \<open>n\<close> and \<open>l\<close> are natural numbers.
+Thus \<open>2*n\<close> can be less than \<open>l\<close> (in which case \<open>2*n - l = 0\<close>!)
+because the invariant does not exclude this case, although it cannot arise.
+If you go through the proof of lemma \<open>a_ins\<close> in detail
+you will understand why this case also works out fine.\<close>
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+apply(auto)
+done
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+apply(cases t)
+apply(auto split: if_splits)
+done
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Kurz_Friedrich_friedrich.kurz@tum.de_606/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.QZEOO6HM4F'
+Files in /tmp/eval-606-j00vOy: hw_12.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Madge_Pimentel_Fabio_fabio.madge@tum.de_594/Amortized_ExamplesX.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,242 @@
+(** Score: 5/5
+ why not simply deleting the irrelevant stuff, such that I can find your solution more easily?
+*)
+
+section "Amortized Complexity: Classical Examples"
+
+theory Amortized_ExamplesX
+imports Complex_Main
+begin
+
+subsection "Binary Counter"
+
+locale Counter
+begin
+
+type_synonym counter = "bool list"
+
+definition init :: counter where
+"init = []"
+
+fun incr :: "counter \<Rightarrow> counter" where
+"incr [] = [True]" |
+"incr (False # bs) = True # bs" |
+"incr (True # bs) = False # incr bs"
+
+fun t_incr :: "counter \<Rightarrow> nat" where
+"t_incr [] = 1" |
+"t_incr (False # bs) = 1" |
+"t_incr (True # bs) = t_incr bs + 1"
+
+definition \<Phi> :: "counter \<Rightarrow> int" where
+"\<Phi> bs = length(filter id bs)"
+
+lemma \<Phi>_non_neg: "\<Phi> bs \<ge> 0"
+by(simp add: \<Phi>_def)
+
+lemma \<Phi>_init: "\<Phi> init = 0"
+by(simp add: \<Phi>_def init_def)
+
+lemma a_incr: "t_incr bs + \<Phi>(incr bs) - \<Phi> bs = 2"
+apply(induction bs rule: incr.induct)
+apply (simp_all add: \<Phi>_def)
+done
+
+(* Application: Doing n increments *)
+
+fun n_increments where
+ "n_increments 0 bs = bs"
+| "n_increments (Suc n) bs = n_increments n (incr bs)"
+
+fun t_n_increments :: "nat \<Rightarrow> counter \<Rightarrow> nat" where
+ "t_n_increments 0 bs = 1"
+| "t_n_increments (Suc n) bs = t_incr bs + t_n_increments n (incr bs)"
+
+(* Generalized lemma with arbitrary start state.
+ Telescoping sum is done implicitly in this proof. *)
+lemma t_n_increments_aux:
+ "t_n_increments n bs = 1 + 2*n + \<Phi> bs - \<Phi> (n_increments n bs)"
+proof (induction n bs rule: n_increments.induct)
+ case (1 bs)
+ then show ?case by simp
+next
+ case (2 n bs)
+ then show ?case
+ using a_incr[of bs] by auto
+qed
+
+(* Estimation for initial state *)
+lemma t_n_increments: "t_n_increments n init \<le> 1+2*n"
+ using
+ t_n_increments_aux[of n init]
+ \<Phi>_non_neg[of "n_increments n init"] (* Exploit that potential cannot be negative *)
+ \<Phi>_init
+ by auto
+
+
+end
+
+
+subsection "Dynamic tables: insert only"
+
+locale Dyn_Tab
+begin
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) = (n \<le> l \<and> l \<le> 2 * n)"
+
+definition init :: tab where
+"init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+apply(cases t)
+ by auto
+
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = 2 * (real n) - real l"
+
+text \<open>Careful: \<open>n\<close> and \<open>l\<close> are natural numbers.
+Thus \<open>2*n\<close> can be less than \<open>l\<close> (in which case \<open>2*n - l = 0\<close>!)
+because the invariant does not exclude this case, although it cannot arise.
+If you go through the proof of lemma \<open>a_ins\<close> in detail
+you will understand why this case also works out fine.\<close>
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+ apply(simp)
+done
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+ apply(cases t)
+ apply(auto split: if_splits)
+ done
+
+end
+
+
+subsection "Queue"
+
+locale Queue
+begin
+
+text \<open>Very simplified model, dequeue returns no result.\<close>
+(* REMARK: Using \<open>unit list\<close> would simplify model even more. *)
+
+type_synonym 'a queue = "'a list * 'a list"
+
+definition init :: "'a queue" where
+"init = ([],[])"
+
+fun enq :: "'a \<Rightarrow> 'a queue \<Rightarrow> 'a queue" where
+"enq x (xs,ys) = (x#xs, ys)"
+
+fun deq :: "'a queue \<Rightarrow> 'a queue" where
+"deq (xs,ys) = (if ys = [] then ([], tl(rev xs)) else (xs, tl ys))"
+
+text \<open>Time: we count only the number of newly allocated list cells.\<close>
+
+fun t_enq :: "'a \<Rightarrow> 'a queue \<Rightarrow> nat" where
+"t_enq x (xs,ys) = 1"
+
+text \<open>We assume that function \<open>rev\<close> has linear complexity.\<close>
+
+fun t_deq :: "'a queue \<Rightarrow> nat" where
+"t_deq (xs,ys) = (if ys = [] then length xs else 0)"
+
+fun \<Phi> :: "'a queue \<Rightarrow> int" where
+"\<Phi> (xs,ys) = length xs"
+
+lemma \<Phi>_non_neg: "\<Phi> q \<ge> 0"
+by(cases q) simp
+
+lemma \<Phi>_init: "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_enq: "t_enq x q + \<Phi>(enq x q) - \<Phi> q = 2"
+by(cases q) auto
+
+lemma a_deq: "t_deq q + \<Phi>(deq q) - \<Phi> q = 0"
+by(cases q) auto
+
+text \<open>
+ Estimating the cost of a sequence of enqueue and dequeue operations.
+ Note: The cost of a dequeue operation on the empty list is just 0, thus
+ that we do not care about the queue being actually non-empty if we
+ call \<open>deq\<close>.
+\<close>
+
+lemma "t_deq init = 0" by (auto simp: init_def)
+
+datatype 'a opr = ENQ 'a | DEQ
+
+fun execute :: "'a queue \<Rightarrow> 'a opr list \<Rightarrow> 'a queue" where
+ "execute q [] = q"
+| "execute q (ENQ x#ops) = execute (enq x q) ops"
+| "execute q (DEQ#ops) = execute (deq q) ops"
+
+fun t_execute :: "'a queue \<Rightarrow> 'a opr list \<Rightarrow> nat" where
+ "t_execute q [] = 0"
+| "t_execute q (ENQ x#ops) = t_enq x q + t_execute (enq x q) ops"
+| "t_execute q (DEQ#ops) = t_deq q + t_execute (deq q) ops"
+
+lemma "t_execute q ops \<le> 2*length ops + \<Phi> q - \<Phi> (execute q ops)"
+ apply (induction q ops rule: execute.induct)
+ apply auto
+ done
+
+(* Note: The simplifier re-proves a_enq and a_deq here, as the definitions of
+ the operations and potential are in the simpset by default.
+
+ This may be nice for simple properties, but for more complex properties,
+ it will just result in very deep unfolding, exploding the subgoals and thus
+ rendering them intractable.
+
+ Thus, we clean up the simpset to not unfold below our abstraction boundary,
+ which is the operations and the potential:
+*)
+lemmas [simp del] = \<Phi>.simps enq.simps deq.simps t_enq.simps t_deq.simps
+
+(* This proof structure will also work for more complicated estimates! *)
+lemma t_execute_aux: "t_execute q ops \<le> 2*length ops + \<Phi> q - \<Phi> (execute q ops)"
+proof (induction q ops rule: execute.induct)
+ case (1 q)
+ then show ?case by simp
+next
+ case (2 q x ops)
+ then show ?case using a_enq[of x q] by simp
+next
+ case (3 q ops)
+ then show ?case using a_deq[of q] by simp
+qed
+
+
+lemma t_execute: "t_execute init ops \<le> 2*length ops"
+ using t_execute_aux[of init ops] \<Phi>_non_neg[of "execute init ops"]
+ by (auto simp: \<Phi>_init)
+
+end
+
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Madge_Pimentel_Fabio_fabio.madge@tum.de_594/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+fabio.madge@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Madge_Pimentel_Fabio_fabio.madge@tum.de_594/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.sJIsrFcOMI'
+Files in /tmp/eval-594-gwyznl: Amortized_Examples.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Ouyang_Lena_ga96vup@mytum.de_598/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+ga96vup@mytum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Ouyang_Lena_ga96vup@mytum.de_598/hw12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,61 @@
+(** Score: 5/5
+*)
+theory hw12
+imports Complex_Main
+begin
+subsection "Dynamic tables: insert only"
+
+locale Dyn_Tab
+begin
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) = ((n \<le> l) \<and> (l \<le> 2*n))"
+
+definition init :: tab where
+"init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+apply(cases t)
+apply(auto)
+done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = 2*(real n) - (real l)"
+
+text \<open>Careful: \<open>n\<close> and \<open>l\<close> are natural numbers.
+Thus \<open>2*n\<close> can be less than \<open>l\<close> (in which case \<open>2*n - l = 0\<close>!)
+because the invariant does not exclude this case, although it cannot arise.
+If you go through the proof of lemma \<open>a_ins\<close> in detail
+you will understand why this case also works out fine.\<close>
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+apply(auto)
+done
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+apply(cases t)
+apply(auto split: if_splits)
+done
+end
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Ouyang_Lena_ga96vup@mytum.de_598/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.vujKdEsBER'
+Files in /tmp/eval-598-DpEGnX: hw12.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Rdle_Karl_Jonas_jonas.raedle@tum.de_605/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+jonas.raedle@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Rdle_Karl_Jonas_jonas.raedle@tum.de_605/hw12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,51 @@
+(** Score: 5/5
+*)
+theory hw12
+imports
+ Complex_Main
+begin
+
+locale Dyn_Tab
+begin
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) \<longleftrightarrow> (n \<le> l) \<and> (l \<le> 2 * n)"
+
+definition init :: tab where
+"init = (0,0)"
+
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+apply(cases t)
+apply(auto)
+done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = 2 * (real n) - (real l)"
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+apply(auto)
+done
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+apply(cases t)
+apply(auto split: if_splits)
+done
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Rdle_Karl_Jonas_jonas.raedle@tum.de_605/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.lw6Bkh4oHQ'
+Files in /tmp/eval-605-guP0Ih: hw12.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Rokopf_Simon_simon.rosskopf@tum.de_601/Amortized_ExamplesX.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,158 @@
+(** Score: 5/5
+ the solution was intended to be easy ... this, unfortunately, makes it a bit boring ;)
+*)
+
+theory Amortized_ExamplesX
+imports Complex_Main
+begin
+
+subsection "Dynamic tables: insert only"
+
+locale Dyn_Tab
+begin
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+type_synonym tab = "nat \<times> nat"
+fun invar :: "tab \<Rightarrow> bool" where
+(* Boring solution, just add what first offending lemma wants to definition ;) *)
+"invar (n,l) = (n \<le> l \<and> (2*(real n) - real l) \<ge> 0)"
+
+definition init :: tab where
+"init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+apply(cases t)
+apply(auto)
+done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = 2*(real n) - real l"
+
+text \<open>Careful: \<open>n\<close> and \<open>l\<close> are natural numbers.
+Thus \<open>2*n\<close> can be less than \<open>l\<close> (in which case \<open>2*n - l = 0\<close>!)
+because the invariant does not exclude this case, although it cannot arise.
+If you go through the proof of lemma \<open>a_ins\<close> in detail
+you will understand why this case also works out fine.\<close>
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+apply(auto)
+done
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+apply(cases t)
+apply(auto split: if_splits)
+done
+
+end
+
+
+subsection "Dynamic tables: insert and delete"
+
+(*
+
+No longer possible to proof lemma del_invar.
+Stuff above requires invariant to be strong enough to show (2*(real n) - real l) \<ge> 0.
+If one can now proof del_invar,
+then invariant must hold for (n,l) = (1,4) = del (del (ins (ins (ins init))))
+but it does not...
+
+locale Dyn_Tab2 = Dyn_Tab
+begin
+
+fun del :: "tab \<Rightarrow> tab" where
+"del (n,l) = (n-1, if n=1 then 0 else if 4*(n-1) < l then l div 2 else l)"
+
+fun t_del :: "tab \<Rightarrow> nat" where
+"t_del (n,l) = (if n=1 then 1 else if 4*(n-1) < l then n else 1)"
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = (if 2*n < l then l/2 - n else 2*n - l)"
+
+lemma \<Phi>_non_negative: "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+by(cases t) auto
+
+lemma \<Phi>_init: "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+by (cases t) (auto split: if_splits)
+
+lemma a_del: "invar t \<Longrightarrow> t_del t + \<Phi>(del t) - \<Phi> t \<le> 2"
+by (cases t) (auto split: if_splits)
+
+lemma init_invar: "invar init" by (auto simp: init_def)
+lemma ins_invar: "invar t \<Longrightarrow> invar (ins t)"
+ by (cases t) (auto)
+
+lemma del_invar: "invar t \<Longrightarrow> invar (del t)"
+ by (cases t) (auto)
+
+text \<open>Establish reasonable abstraction boundary\<close>
+lemmas [simp del] = ins.simps t_ins.simps del.simps t_del.simps \<Phi>.simps
+
+datatype opr = INS | DEL
+
+fun execute :: "tab \<Rightarrow> opr list \<Rightarrow> tab" where
+ "execute t [] = t"
+| "execute t (INS#ops) = execute (ins t) ops"
+| "execute t (DEL#ops) = execute (del t) ops"
+
+fun t_execute :: "tab \<Rightarrow> opr list \<Rightarrow> nat" where
+ "t_execute t [] = 1"
+| "t_execute t (INS#ops) = t_ins t + t_execute (ins t) ops"
+| "t_execute t (DEL#ops) = t_del t + t_execute (del t) ops"
+
+lemma execute_invar: "invar t \<Longrightarrow> invar (execute t ops)"
+ by (induction t ops rule: execute.induct) (auto simp: ins_invar del_invar)
+
+(*
+ To use the amortized complexity estimate here,
+ we have to additionally keep track of the invariant.
+*)
+lemma t_execute_aux: "invar t \<Longrightarrow> t_execute t ops \<le> 1 + 3*length ops + \<Phi> t - \<Phi> (execute t ops)"
+proof (induction t ops rule: execute.induct)
+ case (1 t)
+ then show ?case by simp
+next
+ case (2 t ops)
+ then show ?case
+ using a_ins[of t]
+ by (auto simp: ins_invar)
+
+next
+ case (3 t ops)
+ then show ?case
+ using a_del[of t]
+ by (auto simp: del_invar)
+
+qed
+
+lemma t_execute: "t_execute init ops \<le> 1 + 3*length ops"
+ using
+ t_execute_aux[of init ops]
+ \<Phi>_non_negative[OF execute_invar, of init ops]
+ init_invar
+ \<Phi>_init
+ by auto
+
+end
+*)
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Rokopf_Simon_simon.rosskopf@tum.de_601/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+simon.rosskopf@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Rokopf_Simon_simon.rosskopf@tum.de_601/testoutput.html Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,24 @@
+
+<html>
+<head>
+<link rel="stylesheet" href="https://maxcdn.bootstrapcdn.com/bootstrap/3.3.7/css/bootstrap.min.css">
+</head>
+<body>
+<h3>1/1 test cases passed!</h3>
+
+<table class="table">
+<thead>
+<tr>
+<th>Test Case</th><th>Passed</th>
+</tr>
+</thead>
+<tbody>
+
+<tr><td>check_build</td><td>Passed</td></tr>
+
+</tbody>
+</table>
+
+</body>
+</html>
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Saporta_Antoine_Jacques_antoine.saporta@tum.de_580/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+antoine.saporta@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Saporta_Antoine_Jacques_antoine.saporta@tum.de_580/hw12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,63 @@
+(** Score: 5/5
+*)
+(*<*)
+theory hw12
+imports Complex_Main
+begin
+(*>*)
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) = (n \<le> l \<and> l \<le> 2*n)"
+
+definition init :: tab where
+"init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+apply(cases t)
+apply(auto)
+done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = 2*(real n) - real l"
+
+text \<open>Careful: \<open>n\<close> and \<open>l\<close> are natural numbers.
+Thus \<open>2*n\<close> can be less than \<open>l\<close> (in which case \<open>2*n - l = 0\<close>!)
+because the invariant does not exclude this case, although it cannot arise.
+If you go through the proof of lemma \<open>a_ins\<close> in detail
+you will understand why this case also works out fine.\<close>
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+apply(auto)
+done
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+apply(cases t)
+apply(auto split: if_splits)
+done
+
+
+(*<*)
+end
+(*>*)
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Saporta_Antoine_Jacques_antoine.saporta@tum.de_580/testoutput.html Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,24 @@
+
+<html>
+<head>
+<link rel="stylesheet" href="https://maxcdn.bootstrapcdn.com/bootstrap/3.3.7/css/bootstrap.min.css">
+</head>
+<body>
+<h3>1/1 test cases passed!</h3>
+
+<table class="table">
+<thead>
+<tr>
+<th>Test Case</th><th>Passed</th>
+</tr>
+</thead>
+<tbody>
+
+<tr><td>check_build</td><td>Passed</td></tr>
+
+</tbody>
+</table>
+
+</body>
+</html>
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Schffeler_Maximilian_maximilian.schaeffeler@tum.de_611/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+maximilian.schaeffeler@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Schffeler_Maximilian_maximilian.schaeffeler@tum.de_611/hw12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,61 @@
+(** Score: 5/5
+*)
+theory hw12
+imports Complex_Main
+begin
+
+locale Dyn_Tab
+begin
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) = (n \<le> l \<and> 2*n \<ge> l)"
+
+definition init :: tab where
+"init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+apply(cases t)
+apply(auto)
+done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = 2*(real n) - real l"
+
+text \<open>Careful: \<open>n\<close> and \<open>l\<close> are natural numbers.
+Thus \<open>2*n\<close> can be less than \<open>l\<close> (in which case \<open>2*n - l = 0\<close>!)
+because the invariant does not exclude this case, although it cannot arise.
+If you go through the proof of lemma \<open>a_ins\<close> in detail
+you will understand why this case also works out fine.\<close>
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+apply(auto)
+done
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+apply(cases t)
+apply(auto split: if_splits)
+done
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Schffeler_Maximilian_maximilian.schaeffeler@tum.de_611/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.Ci4f3jckge'
+Files in /tmp/eval-611-XU4W1w: hw12.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Somogyi_Daniel_daniel.somogyi@tum.de_578/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+daniel.somogyi@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Somogyi_Daniel_daniel.somogyi@tum.de_578/hw12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,62 @@
+(** Score: 5/5
+*)
+theory hw12
+ imports Complex_Main
+
+begin
+ subsection "Dynamic tables: insert only"
+
+locale Dyn_Tab
+begin
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) = ((n \<le> l) \<and> (if l = 0 then (n = 0) else (l \<le> 2*n)))"
+
+definition init :: tab where
+"init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma"invar t \<Longrightarrow> invar(ins t)"
+ by(cases t) auto
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = 2*(real n) - l"
+
+text \<open>Careful: \<open>n\<close> and \<open>l\<close> are natural numbers.
+Thus \<open>2*n\<close> can be less than \<open>l\<close> (in which case \<open>2*n - l = 0\<close>!)
+because the invariant does not exclude this case, although it cannot arise.
+If you go through the proof of lemma \<open>a_ins\<close> in detail
+you will understand why this case also works out fine.\<close>
+
+lemma
+ "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+ by(cases t)(auto split:if_splits)
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+apply(cases t)
+apply(auto split: if_splits)
+done
+
+end
+
+end
+
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Somogyi_Daniel_daniel.somogyi@tum.de_578/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.bK1RNA1o7U'
+Files in /tmp/eval-578-p4ZoRZ: hw12.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Stevens_Lukas_lukas.stevens@tum.de_574/Ex12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,62 @@
+(** Score: 5/5
+*)
+theory Ex12
+ imports Complex_Main
+begin
+
+locale Dyn_Tab
+begin
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+ "invar (n,l) = (n \<le> l \<and> 2 * n \<ge> l)"
+
+definition init :: tab where
+ "init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+ "ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+ "t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+ by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+ apply(cases t)
+ apply(auto)
+ done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+ "\<Phi> (n,l) = 2*(real n) - real l"
+
+text \<open>Careful: \<open>n\<close> and \<open>l\<close> are natural numbers.
+Thus \<open>2*n\<close> can be less than \<open>l\<close> (in which case \<open>2*n - l = 0\<close>!)
+because the invariant does not exclude this case, although it cannot arise.
+If you go through the proof of lemma \<open>a_ins\<close> in detail
+you will understand why this case also works out fine.\<close>
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+ apply(cases t)
+ apply(auto)
+ done
+
+lemma "\<Phi> init = 0"
+ by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+ apply(cases t)
+ apply(auto split: if_splits)
+ done
+end
+
+end
+
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Stevens_Lukas_lukas.stevens@tum.de_574/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+lukas.stevens@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Stevens_Lukas_lukas.stevens@tum.de_574/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.KYWCgOsZm2'
+Files in /tmp/eval-574-0x5b0D: Ex12.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Stwe_Daniel_daniel.stuewe@tum.de_603/HW12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,62 @@
+(** Score: 5/5
+*)
+theory HW12
+imports Complex_Main
+begin
+
+subsection "Dynamic tables: insert only"
+
+locale Dyn_Tab
+begin
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) = (n \<le> l \<and> l \<le> n*2)"
+
+definition init :: tab where
+"init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+apply(cases t)
+apply(auto)
+done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = 2 * (real n) - real l"
+
+text \<open>Careful: \<open>n\<close> and \<open>l\<close> are natural numbers.
+Thus \<open>2*n\<close> can be less than \<open>l\<close> (in which case \<open>2*n - l = 0\<close>!)
+because the invariant does not exclude this case, although it cannot arise.
+If you go through the proof of lemma \<open>a_ins\<close> in detail
+you will understand why this case also works out fine.\<close>
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+apply(auto)
+done
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+apply(cases t)
+apply(auto split: )
+done
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Stwe_Daniel_daniel.stuewe@tum.de_603/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+daniel.stuewe@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Stwe_Daniel_daniel.stuewe@tum.de_603/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.7Jrpb8mkcx'
+Files in /tmp/eval-603-LBPOFL: HW12.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Totakura_Sree_Harsha_sreeharsha.totakura@mytum.de_599/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+sreeharsha.totakura@mytum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Totakura_Sree_Harsha_sreeharsha.totakura@mytum.de_599/hw12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,66 @@
+(** Score: 5/5
+*)
+theory hw12
+ imports Complex_Main
+begin
+
+locale Dyn_Tab
+begin
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+
+type_synonym tab = "nat \<times> nat"
+
+text \<open>The invariant needs to be extended with one more clause to specify that the length of the list
+will never be more than twice the number of elements. This is because we only double the length
+once the list is full.\<close>
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) \<longleftrightarrow> (n \<le> l) \<and> (l \<le> 2 * n)"
+
+definition init :: tab where
+"init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma invar_init: "invar init"
+ by (simp add: init_def)
+
+lemma invar_ins: "invar t \<Longrightarrow> invar(ins t)"
+apply(cases t)
+apply(auto)
+done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = 2*(real n) - real l"
+
+text \<open>Careful: \<open>n\<close> and \<open>l\<close> are natural numbers.
+Thus \<open>2*n\<close> can be less than \<open>l\<close> (in which case \<open>2*n - l = 0\<close>!)
+because the invariant does not exclude this case, although it cannot arise.
+If you go through the proof of lemma \<open>a_ins\<close> in detail
+you will understand why this case also works out fine.\<close>
+
+lemma potential_ge_0: "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+apply(auto)
+done
+
+lemma potential_init_zero: "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+ apply(cases t)
+ apply auto
+done
+
+end
+
+end
+
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Totakura_Sree_Harsha_sreeharsha.totakura@mytum.de_599/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.C9lNgx77mu'
+Files in /tmp/eval-599-KbDMfv: hw12.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Wauligmann_Martin_martin.wauligmann@tum.de_593/SENTMAIL Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1 @@
+martin.wauligmann@tum.de
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Wauligmann_Martin_martin.wauligmann@tum.de_593/hw12.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,56 @@
+(** Score: 5/5
+*)
+theory hw12
+imports Complex_Main
+begin
+
+locale Dyn_Tab
+begin
+
+text\<open>Because we are not interested in the elements of the dynamic table but merely its size
+we abstract it to a pair \<open>(n,l)\<close> where \<open>n\<close> is the number of elements in the table and \<open>l\<close>
+the size of the table.\<close>
+
+type_synonym tab = "nat \<times> nat"
+
+fun invar :: "tab \<Rightarrow> bool" where
+"invar (n,l) \<longleftrightarrow> (n \<le> l) \<and> (l \<le> 2*n)"
+
+definition init :: tab where
+"init = (0,0)"
+
+text\<open>Insertion: the element does not matter\<close>
+fun ins :: "tab \<Rightarrow> tab" where
+"ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2*l)"
+
+text\<open>Time: if the table overflows, we count only the time for copying elements\<close>
+fun t_ins :: "nat \<times> nat \<Rightarrow> nat" where
+"t_ins (n,l) = (if n<l then 1 else n+1)"
+
+lemma "invar init"
+by(simp add: init_def)
+
+lemma "invar t \<Longrightarrow> invar(ins t)"
+apply(cases t)
+apply(auto)
+done
+
+fun \<Phi> :: "tab \<Rightarrow> real" where
+"\<Phi> (n,l) = 2*(real n) - real l"
+
+lemma "invar t \<Longrightarrow> \<Phi> t \<ge> 0"
+apply(cases t)
+apply(auto)
+done
+
+lemma "\<Phi> init = 0"
+by(simp add: init_def)
+
+lemma a_ins: "invar t \<Longrightarrow> t_ins t + \<Phi>(ins t) - \<Phi> t \<le> 3"
+apply(cases t)
+apply(auto split: if_splits)
+done
+
+end
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/Wauligmann_Martin_martin.wauligmann@tum.de_593/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,6 @@
+Using temporary directory '/tmp/tmp.hmCTP05ckL'
+Files in /tmp/eval-593-vNTfz9: hw12.thy
+Submission did not check in Isabelle!
+Test case check_build: Failed
+Runner terminated with exit code 1.
+Test execution terminated with exit code 1.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/12/meta.csv Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,20 @@
+"572","jonas.keinholz@tum.de","Keinholz","Jonas","http://vmnipkow3.in.tum.de/web/submissions/572"
+"573","shuwei.hu@tum.de","Hu","Shuwei","http://vmnipkow3.in.tum.de/web/submissions/573"
+"574","lukas.stevens@tum.de","Stevens","Lukas","http://vmnipkow3.in.tum.de/web/submissions/574"
+"578","daniel.somogyi@tum.de","Somogyi","Daniel","http://vmnipkow3.in.tum.de/web/submissions/578"
+"580","antoine.saporta@tum.de","Saporta","Antoine Jacques","http://vmnipkow3.in.tum.de/web/submissions/580"
+"582","max.haslbeck@mytum.de","Haslbeck","Maximilian","http://vmnipkow3.in.tum.de/web/submissions/582"
+"588","julian.biendarra@tum.de","Biendarra","Julian","http://vmnipkow3.in.tum.de/web/submissions/588"
+"593","martin.wauligmann@tum.de","Wauligmann","Martin","http://vmnipkow3.in.tum.de/web/submissions/593"
+"594","fabio.madge@tum.de","Madge Pimentel","Fabio","http://vmnipkow3.in.tum.de/web/submissions/594"
+"597","markus.grosser@tum.de","Großer","Markus","http://vmnipkow3.in.tum.de/web/submissions/597"
+"598","ga96vup@mytum.de","Ouyang","Lena","http://vmnipkow3.in.tum.de/web/submissions/598"
+"599","sreeharsha.totakura@mytum.de","Totakura","Sree Harsha","http://vmnipkow3.in.tum.de/web/submissions/599"
+"601","simon.rosskopf@tum.de","Roßkopf","Simon","http://vmnipkow3.in.tum.de/web/submissions/601"
+"602","ga95luy@mytum.de","Hutzler","Matthias","http://vmnipkow3.in.tum.de/web/submissions/602"
+"603","daniel.stuewe@tum.de","Stüwe","Daniel","http://vmnipkow3.in.tum.de/web/submissions/603"
+"604","ga48kog@mytum.de","Erhard","Julian","http://vmnipkow3.in.tum.de/web/submissions/604"
+"605","jonas.raedle@tum.de","Rädle","Karl Jonas","http://vmnipkow3.in.tum.de/web/submissions/605"
+"606","friedrich.kurz@tum.de","Kurz","Friedrich","http://vmnipkow3.in.tum.de/web/submissions/606"
+"610","f.hellauer@tum.de","Hellauer","Fabian","http://vmnipkow3.in.tum.de/web/submissions/610"
+"611","maximilian.schaeffeler@tum.de","Schäffeler","Maximilian","http://vmnipkow3.in.tum.de/web/submissions/611"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Biendarra_Julian_julian.biendarra@tum.de_632/matroids.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,620 @@
+
+(*\<Rightarrow>Manuel*)
+theory matroids
+ imports Main
+begin
+
+(*
+Theory of independent systems and matroids
+
+based on the material of "Combinatorial Optimization (MA 4502)",
+SoSe 2017, from José Soto
+
+which is based on "Lee: A First Course in Combinatorial Optimization,
+Cambridge Texts in Applied Mathematics, 2004"
+
+The main part is the Greedy Algorithm and the optimality proof for the
+Greedy Algorithm for matroids.
+*)
+
+locale independent_system =
+ fixes E :: "'a set"
+ fixes I :: "'a set set"
+ assumes powerset: "A \<in> I \<Longrightarrow> A \<subseteq> E"
+ and finite_E: "finite E"
+ and I1: "{} \<in> I"
+ and I2: "\<And>X Y. X \<subseteq> Y \<Longrightarrow> Y \<in> I \<Longrightarrow> X \<in> I"
+begin
+
+abbreviation "base X B \<equiv> B \<subseteq> X \<and> B \<in> I \<and> (\<forall>e \<in> X - B. B \<union> {e} \<notin> I)"
+definition "rank X = Max (card ` {Y. Y \<subseteq> X \<and> Y \<in> I})"
+abbreviation "circuit X \<equiv> X \<subseteq> E \<and> X \<notin> I \<and> (\<forall>e\<in>X. X - {e} \<in> I)"
+abbreviation "circuits \<equiv> {X. circuit X}"
+
+(* finite *)
+lemma finite_I: "finite I"
+ using finite_E powerset
+ by (meson Sup_le_iff finite_UnionD finite_subset)
+
+lemma finite_ind_set:
+ assumes "A \<in> I"
+ shows "finite A"
+ using assms finite_E powerset
+ by (simp add: rev_finite_subset)
+
+(* bases *)
+lemma exist_base:
+ assumes "X \<subseteq> E"
+ shows "\<exists>B. base X B"
+proof -
+ from finite_E assms have "finite X" by (simp add: finite_subset)
+ then show ?thesis
+ proof (induction X rule: finite_induct)
+ case empty
+ from I1 show ?case by simp
+ next
+ case (insert x F)
+ then obtain B where B: "base F B" by auto
+ show ?case
+ proof (cases "B \<union> {x} \<in> I")
+ case True
+ from B have "(\<forall>e \<in> F - B. B \<union> {e} \<notin> I)" by simp
+ have "(\<forall>e \<in> F - (B \<union> {x}). B \<union> {x} \<union> {e} \<notin> I)"
+ proof
+ fix e
+ assume "e \<in> F - (B \<union> {x})"
+ then have "e \<in> F - B" by simp
+ with B have "B \<union> {e} \<notin> I" by simp
+ with I2[of "B \<union> {e}" "B \<union> {x} \<union> {e}"] show "B \<union> {x} \<union> {e} \<notin> I" by auto
+ qed
+ with B True have "base (insert x F) (insert x B)"
+ by auto
+ then show ?thesis by blast
+ next
+ case False
+ with B show ?thesis by auto
+ qed
+ qed
+qed
+
+(* circuits *)
+lemma C1:
+ shows "{} \<notin> circuits"
+by (simp add: I1)
+
+lemma C2:
+ assumes "X \<in> circuits"
+ assumes "Y \<in> circuits"
+ assumes "X \<subseteq> Y"
+ shows "X = Y"
+proof (rule ccontr)
+ assume "X \<noteq> Y"
+ with assms(3) obtain e where e: "e \<in> Y - X" by blast
+ with assms(2) have ind: "Y - {e} \<in> I" by simp
+ with e assms(3) have "X \<subseteq> Y - {e}" by auto
+ with ind assms(1) I2 show False by simp
+qed
+
+(* rank *)
+lemma R1:
+ assumes "finite X"
+ shows "rank X \<in> {0..card X}"
+proof -
+ from finite_I have finite: "finite (card ` {Y. Y \<subseteq> X \<and> Y \<in> I})" by simp
+ have "card X \<ge> rank X"
+ proof (rule ccontr)
+ assume "\<not> rank X \<le> card X"
+ then have asm: "rank X > card X" by simp
+ from I1 have "card {} \<in> card ` {Y. Y \<subseteq> X \<and> Y \<in> I}" by force
+ with Max_in[OF finite] have "Max (card ` {Y. Y \<subseteq> X \<and> Y \<in> I}) \<in> card ` {Y. Y \<subseteq> X \<and> Y \<in> I}"
+ unfolding rank_def
+ by auto
+ with asm have "\<exists>Y \<subseteq> X. Y \<in> I \<and> card Y > card X"
+ unfolding rank_def
+ by auto
+ then obtain Y where "Y \<subseteq> X \<and> card Y > card X" by auto
+ with card_mono[OF assms, of Y] show False by simp
+ qed
+ then show ?thesis by simp
+qed
+
+lemma R2:
+ assumes "X \<subseteq> Y"
+ shows "rank X \<le> rank Y"
+proof -
+ from assms have *: "card ` {Z. Z \<subseteq> X \<and> Z \<in> I} \<subseteq> card ` {Z. Z \<subseteq> Y \<and> Z \<in> I}"
+ by auto
+ from finite_I have "finite (card ` {Z. Z \<subseteq> Y \<and> Z \<in> I})"
+ by simp
+ with Max_mono[OF *] I1 show ?thesis
+ unfolding rank_def
+ by auto
+qed
+
+lemma rank_ind:
+ assumes "X \<in> I"
+ shows "rank X = card X"
+proof -
+ from finite_ind_set assms have "finite X" by simp
+ from finite_I have finite: "finite (card ` {Y. Y \<subseteq> X \<and> Y \<in> I})" by simp
+ also from assms have "card X \<in> card ` {Y. Y \<subseteq> X \<and> Y \<in> I}" by simp
+ ultimately have "card X \<le> rank X" unfolding rank_def by auto
+ also from R1[OF `finite X`] have "rank X \<le> card X" by simp
+ finally show ?thesis ..
+qed
+
+end
+
+
+locale matroid =
+ fixes E :: "'a set"
+ fixes I :: "'a set set"
+ assumes independent: "independent_system E I"
+ and I3: "\<And>X Y. X \<in> I \<Longrightarrow> Y \<in> I \<Longrightarrow> card X < card Y \<Longrightarrow> \<exists>e \<in> Y - X. X \<union> {e} \<in> I"
+begin
+
+sublocale independent_system by (rule independent)
+
+lemma I3':
+ assumes "X \<in> I"
+ assumes "Y \<in> I"
+ assumes "card X \<le> card Y"
+ shows "\<exists>Z. X \<subseteq> Z \<and> Z \<in> I \<and> card Y = card Z"
+proof -
+ from finite_ind_set assms(2) have "finite Y" by simp
+ from this assms show ?thesis
+ proof (induction Y arbitrary: X rule: finite_induct)
+ case empty
+ with I1 finite_ind_set show ?case by auto
+ next
+ case (insert x F)
+ from I2[OF _ insert.prems(2), of F] have "F \<in> I"
+ by (simp add: subset_insertI)
+ show ?case
+ proof (cases "card X = card (insert x F)")
+ case True
+ with `X \<in> I` have "X \<subseteq> X \<and> X \<in> I \<and> card (insert x F) = card X"
+ by simp
+ then show ?thesis by auto
+ next
+ case False
+ with `card X \<le> card (insert x F)` `x \<notin> F` `finite F`
+ have "card X \<le> card F" by simp
+ from insert.IH[OF `X \<in> I` `F \<in> I` this]
+ obtain Z where "X \<subseteq> Z" "Z \<in> I" and "card F = card Z"
+ by auto
+ with `F \<in> I` finite_ind_set `x \<notin> F` have "card Z + 1 = card (insert x F)"
+ by simp
+ then have "card Z < card (insert x F)" by simp
+ from I3[OF `Z \<in> I` `insert x F \<in> I` this]
+ obtain e where "e \<in> insert x F - Z" and "Z \<union> {e} \<in> I"
+ by auto
+ with `card Z + 1 = card (insert x F)` `X \<subseteq> Z` `Z \<in> I` finite_ind_set
+ have "X \<subseteq> insert e Z \<and> insert e Z \<in> I \<and> card (insert x F) = card (insert e Z)"
+ by auto
+ then show ?thesis by auto
+ qed
+ qed
+qed
+
+(* base *)
+lemma
+ assumes "base X A"
+ assumes "base X B"
+ shows "card A = card B"
+proof (rule ccontr)
+ assume asm: "card A \<noteq> card B"
+ show False
+ proof (cases "card A < card B")
+ case True
+ from assms(1) have "A \<in> I" by simp
+ also from assms(2) have "B \<in> I" by simp
+ ultimately obtain e where "e \<in> B - A" and "A \<union> {e} \<in> I" using I3 True by blast
+ with assms(2) have "e \<in> X - A" by auto
+ with assms(1) have "A \<union> {e} \<notin> I" by simp
+ with `A \<union> {e} \<in> I` show ?thesis by simp
+ next
+ case False
+ with asm have False': "card B < card A" by simp
+ from assms(1) have "A \<in> I" by simp
+ also from assms(2) have "B \<in> I" by simp
+ ultimately obtain e where "e \<in> A - B" and "B \<union> {e} \<in> I" using I3 False' by blast
+ with assms(1) have "e \<in> X - B" by auto
+ with assms(2) have "B \<union> {e} \<notin> I" by simp
+ with `B \<union> {e} \<in> I` show ?thesis by simp
+ qed
+qed
+
+(* rank *)
+lemma rank_base:
+ assumes "base X B"
+ shows "rank X = card B"
+proof -
+ from finite_ind_set assms have "finite B" by simp
+ from finite_I have finite: "finite (card ` {Y. Y \<subseteq> X \<and> Y \<in> I})" by simp
+ also from assms have B_in: "card B \<in> card ` {Y. Y \<subseteq> X \<and> Y \<in> I}" by simp
+ ultimately have "card B \<le> rank X" unfolding rank_def by auto
+ also have "rank X \<le> card B"
+ proof (rule ccontr)
+ assume asm: "\<not> rank X \<le> card B"
+ from Max_in[OF finite] B_in have "rank X \<in> card ` {Y. Y \<subseteq> X \<and> Y \<in> I}"
+ unfolding rank_def by auto
+ then obtain Y where "card Y = rank X" and "Y \<subseteq> X" and "Y \<in> I"
+ by auto
+ with asm have "card Y > card B" by simp
+ with I3 assms `Y \<in> I` obtain e where "e \<in> Y - B" and "B \<union> {e} \<in> I"
+ by blast
+ with `Y \<subseteq> X` have "e \<in> X - B" by auto
+ with `B \<union> {e} \<in> I` assms show False by simp
+ qed
+ finally show ?thesis ..
+qed
+
+(* Greedy *)
+fun greedy_s :: "'a list \<Rightarrow> 'a list" where
+ "greedy_s [] = []" |
+ "greedy_s (e#es) = (
+ if set (e # greedy_s es) \<in> I
+ then e # greedy_s es
+ else greedy_s es)"
+
+fun greedy :: "('a \<Rightarrow> nat) \<Rightarrow> 'a set" where
+ "greedy w = set (greedy_s (linorder.sorted_list_of_set (\<lambda>a b. w a \<le> w b) E))"
+
+lemma greedy_property:
+ assumes "i < length es"
+ assumes "es ! i \<notin> set (greedy_s es)"
+ shows "set (greedy_s es) \<union> {es ! i} \<notin> I"
+ using assms
+ proof (induction es arbitrary: i rule: greedy_s.induct)
+ case 1
+ then show ?case by simp
+ next
+ case (2 e es)
+ show ?case
+ proof(cases i)
+ case 0
+ with 2 show ?thesis by auto
+ next
+ case (Suc nat)
+ with 2(1)[of nat] 2(4) 2(5)
+ I2[of "insert (es ! nat) (set (greedy_s es))"
+ "insert (es ! nat) (insert e (set (greedy_s es)))"]
+ subset_insertI2[OF subset_insertI[of "set (greedy_s es)" "e"]]
+ show ?thesis
+ by auto
+ qed
+ qed
+
+lemma greedy_s_sublist:
+ shows "greedy_s es \<in> set (sublists es)"
+proof (induction es)
+ case Nil
+ then show ?case by simp
+next
+ case (Cons e es)
+ have "sublists (e#es) = map (op # e) (sublists es) @ (sublists es)"
+ by (meson sublists.simps(2))
+ with Cons.IH show ?case
+ by auto
+qed
+
+lemma greedy_s_subset:
+ shows "set (greedy_s es) \<subseteq> set es"
+ by (induction es) auto
+
+lemma greedy_s_distinct:
+ assumes "distinct es"
+ shows "distinct (greedy_s es)"
+ using assms greedy_s_subset
+ by (induction es) auto
+
+lemma greedy_s_sorted:
+ assumes "class.linorder leq le"
+ assumes "linorder.sorted leq es"
+ shows "linorder.sorted leq (greedy_s es)"
+ using assms
+ proof (induction es rule: greedy_s.induct)
+ case 1
+ then show ?case by simp
+ next
+ case (2 e es)
+ then show ?case
+ proof (cases "set (e # greedy_s es) \<in> I")
+ case True
+ from linorder.sorted_Cons[OF 2(4)] 2(5)
+ have *: "linorder.sorted leq es \<and> (\<forall>y\<in>set es. leq e y)"
+ by simp
+ with 2(1)[OF 2(4)] have **: "linorder.sorted leq (greedy_s es)"
+ by simp
+ from * greedy_s_subset have "\<forall>y\<in>set (greedy_s es). leq e y" by auto
+ with linorder.sorted_Cons[OF 2(4), of e "greedy_s es"] ** show ?thesis
+ by auto
+ next
+ case False
+ with 2 linorder.sorted_Cons[OF 2(4)] show ?thesis by auto
+ qed
+ qed
+
+lemma greedy_s_take:
+ shows "set (greedy_s (drop k es)) \<subseteq> set (greedy_s es)"
+proof (induction es arbitrary: k)
+ case (Cons a es)
+ then show ?case
+ by (cases k) auto
+qed simp
+
+lemma greedy_s_result:
+ assumes "distinct es"
+ assumes "k < length es"
+ assumes "p < length (greedy_s es)"
+ assumes "es ! k = greedy_s es ! p"
+ shows "greedy_s (drop (k+1) es) = drop (p+1) (greedy_s es)"
+ using assms
+proof (induction es arbitrary: k p)
+ case (Cons a es)
+ from greedy_s_subset[of es]
+ have in_set: "i < length (greedy_s es) \<Longrightarrow> greedy_s es ! i \<in> set es" for i
+ by auto
+ from Cons show ?case
+ proof (cases k)
+ case 0
+ with Cons in_set show ?thesis by (cases p) auto
+ next
+ case (Suc nat)
+ with Cons in_set show ?thesis by (cases p) auto
+ qed
+qed simp
+
+lemma greedy_sol_independent:
+ fixes w :: "'a \<Rightarrow> nat"
+ shows "greedy w \<in> I"
+proof -
+ define es where es: "es = linorder.sorted_list_of_set (\<lambda>a b. w a \<le> w b) E"
+ have "set (greedy_s es) \<in> I"
+ by (induction es rule: greedy_s.induct) (auto simp: I1)
+ with es show ?thesis by simp
+qed
+
+lemma greedy_sol_base:
+ fixes w :: "'a \<Rightarrow> nat"
+ assumes "inj w"
+ shows "base E (greedy w)"
+proof -
+ from assms(1)
+ have linorder: "class.linorder (\<lambda>a b. w a \<le> w b) (\<lambda>a b. w a < w b)"
+ unfolding class.linorder_def class.order_def class.linorder_axioms_def
+ class.preorder_def class.order_axioms_def
+ by (auto simp: injD)
+
+ define es where es: "es = linorder.sorted_list_of_set (\<lambda>a b. w a \<le> w b) E"
+ have "\<forall>e\<in>E - greedy w. greedy w \<union> {e} \<notin> I"
+ proof
+ fix e
+ assume asm: "e \<in> E - greedy w"
+ with es linorder.sorted_list_of_set[OF linorder finite_E]
+ have "e \<in> set es" by simp
+ then obtain i where "es ! i = e" and i_le: "i < length es"
+ by (auto simp: in_set_conv_nth)
+ with asm es have "es ! i \<notin> set (greedy_s es)" by simp
+ from greedy_property[OF i_le this] es `es ! i = e`
+ show "greedy w \<union> {e} \<notin> I" by simp
+ qed
+ with greedy_sol_independent powerset show ?thesis by simp
+qed
+
+theorem greedy_local_optimal:
+ fixes w :: "'a \<Rightarrow> nat"
+ (* further TODO: get rid of constraint `ind w` *)
+ assumes "inj w"
+ assumes "T \<in> I"
+ assumes "card T = card (greedy w)"
+ shows "sum w (greedy w) \<ge> sum w T"
+proof (rule ccontr)
+ assume asm: "\<not> (sum w (greedy w) \<ge> sum w T)"
+ define S where S: "S = greedy w"
+ define k where k: "k = card S"
+ define ss where ss: "ss = linorder.sorted_list_of_set (\<lambda>a b. w a \<le> w b) S"
+ define ts where ts: "ts = linorder.sorted_list_of_set (\<lambda>a b. w a \<le> w b) T"
+ define es where es: "es = linorder.sorted_list_of_set (\<lambda>a b. w a \<le> w b) E"
+ from assms(1)
+ have linorder: "class.linorder (\<lambda>a b. w a \<le> w b) (\<lambda>a b. w a < w b)"
+ unfolding class.linorder_def class.order_def class.linorder_axioms_def
+ class.preorder_def class.order_axioms_def
+ by (auto simp: injD)
+ from S greedy_sol_independent have S_ind: "S \<in> I" by simp
+ with finite_ind_set have finite_S: "finite S"
+ by simp
+ from ss linorder.sorted_list_of_set[OF linorder finite_S]
+ have sorted_ss: "linorder.sorted (\<lambda>a b. w a \<le> w b) ss"
+ and distinct_ss: "distinct ss"
+ and ss_S: "set ss = S"
+ by auto
+ from assms finite_ind_set have finite_T: "finite T"
+ by simp
+ from ts linorder.sorted_list_of_set[OF linorder finite_T]
+ have sorted_ts: "linorder.sorted (\<lambda>a b. w a \<le> w b) ts"
+ and distinct_ts: "distinct ts"
+ and ts_T: "set ts = T"
+ by auto
+ from es linorder.sorted_list_of_set[OF linorder finite_E]
+ have sorted_es: "linorder.sorted (\<lambda>a b. w a \<le> w b) es"
+ and distinct_es: "distinct es"
+ and es_E: "set es = E"
+ by auto
+ from k distinct_card distinct_ss ss_S have k_ss: "k = length ss" by auto
+ from k distinct_card distinct_ts ts_T assms S have k_ts: "k = length ts" by fastforce
+
+ from es S have "set (greedy_s es) = S" by simp
+ with ss linorder.sorted_list_of_set_sort_remdups[OF linorder, of "greedy_s es"]
+ have "ss = linorder.sort_key (\<lambda>a b. w a \<le> w b) (\<lambda>x. x) (remdups (greedy_s es))"
+ by simp
+ also from greedy_s_distinct[OF distinct_es] distinct_remdups_id[of "greedy_s es"]
+ have "\<dots> = linorder.sort_key (\<lambda>a b. w a \<le> w b) (\<lambda>x. x) (greedy_s es)"
+ by presburger
+ also from greedy_s_sorted[OF linorder sorted_es]
+ linorder.sorted_sort_id[OF linorder, of "greedy_s es"]
+ have "\<dots> = greedy_s es"
+ by simp
+ finally have "ss = greedy_s es" .
+
+ have "\<exists>p \<in> {0..<k}. w (ts ! p) > w (ss ! p)"
+ proof (rule ccontr)
+ assume "\<not>(\<exists>p \<in> {0..<k}. w (ts ! p) > w (ss ! p))"
+ then have asm2: "\<forall>p \<in> {0..<k}. w (ts ! p) \<le> w (ss ! p)" by auto
+ from distinct_ts ts_T sum_code[of w ts] distinct_remdups_id[of ts]
+ have "sum w T = sum_list (map w ts)"
+ by simp
+ also have "\<dots> = (\<Sum>p\<in>{0..<k}. w (ts ! p))"
+ by (auto simp: sum_list_sum_nth k_ts)
+ also from sum_mono[of "{0..<k}" "\<lambda>p. w (ts ! p)" "\<lambda>p. w (ss ! p)"] asm2
+ have "\<dots> \<le> (\<Sum>p\<in>{0..<k}. w (ss ! p))"
+ by simp
+ also have "\<dots> = sum_list (map w ss)"
+ by (auto simp: sum_list_sum_nth k_ss)
+ also from distinct_ss ss_S sum_code[of w ss] distinct_remdups_id[of ss]
+ have "\<dots> = sum w S"
+ by simp
+ finally show False using S asm by simp
+ qed
+ then obtain p where p_ind: "p\<in>{0..<k}" and w_p: "w (ss ! p) < w (ts ! p)" by blast
+ define T' where T': "T' = set (drop p ts)"
+ define S' where S': "S' = set (drop (p+1) ss)"
+
+ from T' S' distinct_ts distinct_ss p_ind k_ss k_ts have "card S' < card T'"
+ by (auto simp: distinct_card)
+ have "S' \<subseteq> S" using ss_S S' set_drop_subset[of _ ss]
+ by simp
+ with I2 S_ind have "S' \<in> I"
+ by simp
+ have "T' \<subseteq> T" using ts_T T' set_drop_subset[of _ ts]
+ by simp
+ with I2 assms have "T' \<in> I"
+ by simp
+ with `card S' < card T'` `S' \<in> I` I3
+ obtain t where t: "t\<in>T' - S'" and t_ind: "S' \<union> {t} \<in> I"
+ by blast
+ from t T' in_set_conv_nth[of t "drop p ts"]
+ obtain i where "drop p ts ! i = t" and "i<length (drop p ts)"
+ by blast
+ with linorder.sorted_nth_mono[OF linorder linorder.sorted_drop[OF
+ linorder sorted_ts, of p], of 0 i] p_ind k_ts
+ have "w (ts ! p) \<le> w t"
+ by auto
+ with w_p have "w (ss ! p) < w t"
+ by simp
+
+ from ss_S powerset S_ind es_E nth_mem p_ind k_ss
+ have "ss ! p \<in> set es"
+ by auto
+ then obtain j where "es ! j = ss ! p" and "j < length es"
+ by (auto simp: in_set_conv_nth)
+
+ from set_drop_subset[of p ts] t T' ts_T assms(2) powerset es_E
+ have "t \<in> set es"
+ by auto
+ then obtain l where "es ! l = t" and "l < length es"
+ by (auto simp: in_set_conv_nth)
+ from linorder.sorted_nth_mono[OF linorder sorted_es _ `j < length es`,
+ of l] `es ! j = ss ! p` `es ! l = t` `w t > w (ss ! p)`
+ have "j < l"
+ by fastforce
+
+ from `j < l` `l < length es` have l: "l-j-1 < length (drop (j + 1) es)"
+ by simp
+ from `l < length es` `es ! l = t` `j < l` have t': "drop (j+1) es ! (l-j-1) = t"
+ by simp
+ from greedy_s_result[of es j p, OF distinct_es] `es ! j = ss ! p`
+ `ss = greedy_s es` p_ind k_ss `j < length es`
+ have greedy_res: "greedy_s (drop (j + 1) es) = drop (p + 1) ss"
+ by simp
+ with t' S' `t \<in> T' - S'`
+ have "drop (j+1) es ! (l-j-1) \<notin> set (greedy_s (drop (j+1) es))"
+ by simp
+ from S' greedy_res greedy_property[OF l this] t'
+ have "S' \<union> {t} \<notin> I"
+ by simp
+ with `S' \<union> {t} \<in> I` show False by simp
+qed
+
+theorem greedy_optimal:
+ fixes w :: "'a \<Rightarrow> nat"
+ assumes "inj w"
+ shows "greedy w \<in> I \<and> (\<forall>T \<in> I. sum w (greedy w) \<ge> sum w T)"
+proof -
+ from greedy_sol_independent have ind: "greedy w \<in> I" by simp
+ also have "\<forall>T \<in> I. sum w (greedy w) \<ge> sum w T"
+ proof
+ fix T
+ assume "T \<in> I"
+ show "sum w (greedy w) \<ge> sum w T"
+ proof (cases "card T \<le> card (greedy w)")
+ case True
+ from I3'[OF `T \<in> I` ind True]
+ obtain T' where T': "T \<subseteq> T' \<and> T' \<in> I \<and> card (greedy w) = card T'"
+ by blast
+ with finite_ind_set[of T'] have "sum w T \<le> sum w T'"
+ by (auto simp: sum_mono2)
+ also from T' greedy_local_optimal[OF assms] have "sum w T' \<le> sum w (greedy w)"
+ by simp
+ finally show ?thesis .
+ next
+ case False
+ then have False': "card (greedy w) < card T" by simp
+ with I3[OF ind `T \<in> I` False'] obtain e where "e \<in> T - greedy w" and "greedy w \<union> {e} \<in> I"
+ by blast
+ with `T \<in> I` powerset have "e \<in> E - greedy w" by auto
+ from greedy_sol_base[OF assms(1)] have "base E (greedy w)" by simp
+ with `e \<in> E - greedy w` have "greedy w \<union> {e} \<notin> I" by simp
+ with `greedy w \<union> {e} \<in> I` have False by simp
+ then show ?thesis by simp
+ qed
+ qed
+ ultimately show ?thesis ..
+qed
+
+end
+
+(* example: uniform matroid *)
+lemma uniform_matroid:
+ fixes E :: "'a set"
+ fixes r :: nat
+ assumes "finite E"
+ shows "matroid E {X. X \<subseteq> E \<and> card X \<le> r}"
+proof
+ from assms show "finite E" .
+next
+ fix X Y :: "'a set"
+ assume 1: "X \<subseteq> Y"
+ assume 2: "Y \<in> {X. X \<subseteq> E \<and> card X \<le> r}"
+ then have 3: "Y \<subseteq> E" by simp
+ from 2 have 4: "card Y \<le> r" by simp
+ from card_mono[OF finite_subset[OF 3 assms] 1] 1 3 4
+ have "X \<subseteq> E \<and> card X \<le> r" by auto
+ then show "X \<in> {X. X \<subseteq> E \<and> card X \<le> r}" by simp
+next
+ fix X Y :: "'a set"
+ assume 1: "X \<in> {X. X \<subseteq> E \<and> card X \<le> r}"
+ assume 2: "Y \<in> {X. X \<subseteq> E \<and> card X \<le> r}"
+ assume 3: "card X < card Y"
+ from 3 have "Y - X \<noteq> {}"
+ by (metis "1" assms card_empty diff_card_le_card_Diff finite_subset mem_Collect_eq not_le zero_less_diff)
+ then obtain e where e: "e \<in> Y - X" by blast
+ with 1 2 have subset: "X \<union> {e} \<subseteq> E" by auto
+ from e finite_subset[OF subset assms]
+ have "card (X \<union> {e}) = card X + 1"
+ by (auto simp: card_insert_if)
+ also from 3 have "card X + 1 \<le> card Y" by simp
+ also from 2 have "card Y \<le> r" by simp
+ finally have "X \<union> {e} \<in> {X. X \<subseteq> E \<and> card X \<le> r}"
+ using subset by simp
+ with e show "\<exists>e\<in>Y - X. X \<union> {e} \<in> {X. X \<subseteq> E \<and> card X \<le> r}"
+ by auto
+qed auto
+
+interpretation uniform_matroid_ex: matroid "{1::nat..10}" "{X. X \<subseteq> {1..10} \<and> card X \<le> 5}"
+ by (rule uniform_matroid) simp
+
+term uniform_matroid_ex.greedy
+thm uniform_matroid_ex.greedy_optimal
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Biendarra_Julian_julian.biendarra@tum.de_632/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,2 @@
+Using temporary directory '/tmp/tmp.Gsk4eN0mmX'
+Exercise path '/home/admin/suites/exbeoriginal' not found
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Erhard_Julian_ga48kog@mytum.de_655/quadtrees.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,262 @@
+(*\<Rightarrow> Tobias*)
+theory quadtrees
+imports Complex_Main
+begin
+
+type_synonym Point = "real * real"
+
+definition
+ mapping :: "('a \<rightharpoonup> 'b) \<Rightarrow> ('a * 'b) set" where
+ "mapping m = {(a, b). m a = Some b}"
+
+(*Implements a point quadtree: mapping from points to values *)
+(*No duplicate keys are stored*)
+(*if the y value of the new node is strictly greater than of the root of the exisiting tree, it is inserted in the north, ow in the south*)
+(*if the x value of the new node is strictly greater than of the root of the exisiting tree, it is inserted in the east, ow in the west*)
+datatype 'a quadtree = Leaf |
+ Node (key: Point) (val: 'a) (NW: "'a quadtree") (NE: "'a quadtree")
+ (SE: "'a quadtree") (SW: "'a quadtree")
+
+fun qinsert :: "Point \<Rightarrow> 'a \<Rightarrow> 'a quadtree \<Rightarrow> 'a quadtree" where
+ "qinsert (x,y) v Leaf = Node (x,y) v Leaf Leaf Leaf Leaf" |
+ "qinsert (x,y) v (Node (x2, y2) v2 nw ne se sw) = (if x=x2\<and>y=y2 then Node (x2, y2) v nw ne se sw else
+ (if (y > y2) then (if (x \<le> x2) then Node (x2, y2) v2 (qinsert (x,y) v nw) ne se sw
+ else Node (x2, y2) v2 nw (qinsert (x,y) v ne) se sw) else
+ if ( x > x2) then Node (x2, y2) v2 nw ne (qinsert (x,y) v se) sw
+ else Node (x2, y2) v2 nw ne se (qinsert (x,y) v sw)))"
+
+fun qinsert_list :: "'a quadtree \<Rightarrow>(Point * 'a) list \<Rightarrow> 'a quadtree" where
+ "qinsert_list t l = fold (\<lambda>(k,v) acc. qinsert k v acc) l t"
+
+fun list_quadtree :: "'a quadtree \<Rightarrow> (Point * 'a) list" where
+ "list_quadtree Leaf = []" |
+ "list_quadtree (Node (x2, y2) v2 nw ne se sw) = ((x2,y2), v2) # ((list_quadtree nw) @ (list_quadtree ne) @ (list_quadtree se) @ (list_quadtree sw))"
+
+
+fun map_quadtree :: "'a quadtree \<Rightarrow> (Point, 'a) map" where
+ "map_quadtree Leaf = empty" |
+ "map_quadtree (Node (x2, y2) v2 nw ne se sw) = (\<lambda>(x3, y3). if(x3=x2 \<and>y3=y2) then Some v2 else None) ++ ((map_quadtree nw) ++ (map_quadtree ne) ++ (map_quadtree se) ++ (map_quadtree sw))"
+
+fun mapset_quadtree :: "'a quadtree \<Rightarrow> (Point * 'a) set" where
+ "mapset_quadtree Leaf = {}" |
+ "mapset_quadtree (Node (x2, y2) v2 nw ne se sw) = insert ((x2,y2),v2) ((mapset_quadtree nw) \<union> (mapset_quadtree ne) \<union> (mapset_quadtree se) \<union> (mapset_quadtree sw))"
+
+
+fun set_quadtree :: "'a quadtree \<Rightarrow> 'a set" where
+ "set_quadtree Leaf = {}" |
+ "set_quadtree (Node _ v2 nw ne se sw) = insert v2 ((set_quadtree nw) \<union> (set_quadtree ne) \<union> (set_quadtree se) \<union> (set_quadtree sw))"
+
+fun keyset_quadtree :: "'a quadtree \<Rightarrow> Point set" where
+ "keyset_quadtree Leaf = {}" |
+ "keyset_quadtree (Node k _ nw ne se sw) = insert k ((keyset_quadtree nw) \<union> (keyset_quadtree ne) \<union> (keyset_quadtree se) \<union> (keyset_quadtree sw))"
+
+
+fun qlookup :: "Point \<Rightarrow> 'a quadtree \<Rightarrow> 'a option" where
+ "qlookup _ Leaf = None" |
+ "qlookup (x, y) (Node (x2, y2) v2 nw ne se sw) = (if x=x2 \<and> y=y2 then Some v2 else
+ if ( y > y2) then
+ if (x\<le>x2) then (qlookup (x,y) nw)
+ else (qlookup (x,y) ne)
+ else if(x>x2) then (qlookup (x,y) se)
+ else (qlookup (x,y) sw)) "
+
+fun n :: "Point \<Rightarrow> Point \<Rightarrow> bool" where
+"n (_,y) (_,y2) = (y > y2)"
+
+fun e :: "Point \<Rightarrow> Point \<Rightarrow> bool" where
+"e (x,_) (x2,_) = (x > x2)"
+
+fun s :: "Point \<Rightarrow> Point \<Rightarrow> bool" where
+ "s p p2 = ( \<not> n p p2)"
+
+fun w :: "Point \<Rightarrow> Point \<Rightarrow> bool" where
+ "w p p2 = ( \<not> e p p2)"
+
+(*checks whether first arg is nw of second*)
+fun isnw :: "Point \<Rightarrow> Point \<Rightarrow> bool" where
+"isnw p p2 = (n p p2 \<and> w p p2)"
+
+fun isne :: "Point \<Rightarrow> Point \<Rightarrow> bool" where
+"isne p p2 = (n p p2 \<and> e p p2)"
+
+fun issw :: "Point \<Rightarrow> Point \<Rightarrow> bool" where
+"issw p p2 = (s p p2 \<and> w p p2 \<and> p \<noteq> p2)"
+
+fun isse :: "Point \<Rightarrow> Point \<Rightarrow> bool" where
+"isse p p2 = (s p p2 \<and> e p p2)"
+
+
+(* checks whether the all nodes in the subtree of the first argument are in the
+ north west of the point given as the second argument*)
+fun nwOf :: "'a quadtree \<Rightarrow> Point \<Rightarrow> bool" where
+ "nwOf Leaf _ = True" |
+ "nwOf node p2 = (\<forall>p \<in> (keyset_quadtree node) . isnw p p2 )"
+
+fun neOf :: "'a quadtree \<Rightarrow> Point \<Rightarrow> bool" where
+ "neOf Leaf _ = True" |
+ "neOf node p2 = (\<forall>p \<in> (keyset_quadtree node) . isne p p2)"
+
+fun seOf :: "'a quadtree \<Rightarrow> Point \<Rightarrow> bool" where
+ "seOf Leaf _ = True" |
+ "seOf node p2 = (\<forall>p \<in> (keyset_quadtree node) . isse p p2)"
+
+fun swOf :: "'a quadtree \<Rightarrow> Point \<Rightarrow> bool" where
+ "swOf Leaf _ = True" |
+ "swOf node p2 = (\<forall>p \<in> (keyset_quadtree node) . issw p p2)"
+(*end of helper functions*)
+
+fun invar :: "'a quadtree \<Rightarrow> bool" where
+ "invar Leaf = True" |
+ "invar (Node p _ nw ne se sw) \<longleftrightarrow> nwOf nw p \<and> neOf ne p \<and> seOf se p \<and> swOf sw p \<and>
+ invar nw \<and> invar ne \<and> invar se \<and> invar sw"
+
+fun delete :: "Point \<Rightarrow> 'a quadtree \<Rightarrow> 'a quadtree" where
+ "delete _ Leaf = Leaf" |
+ "delete (x,y) (Node (x2, y2) v nw ne se sw) = (if x=x2 \<and> y=y2 then qinsert_list Leaf (tl (list_quadtree (Node (x2, y2) v nw ne se sw))) else
+ (if y > y2 then (if x < x2 then (Node (x2,y2) v (delete (x,y) nw) ne se sw)
+ else (Node (x2, y2) v nw (delete (x,y) ne) se sw ))
+ else (if x \<le> x2 then (Node (x2,y2) v nw ne sw (delete (x,y) sw))
+ else (Node (x2, y2) v nw ne (delete (x,y) se) sw ))))"
+
+lemma ks: " keyset_quadtree(qinsert (x,y) v t) = keyset_quadtree(t) \<union> {(x,y)}"
+ apply (induction t)
+ apply (auto split: if_splits)
+ done
+
+lemma qinsert_non_Leaf: "qinsert (x,y) v t \<noteq> Leaf"
+ apply (cases t)
+ apply (auto split: if_splits)
+ done
+
+lemma "invar t \<Longrightarrow> invar (qinsert (x,y) v t)"
+proof (induction t rule:qinsert.induct)
+ case (1 x y v)
+ then show ?case by simp
+next
+ case (2 x y v x2 y2 v2 nw ne se sw)
+ have "x=x2 \<and> y=y2 \<or> \<not>(x=x2 \<and> y=y2)" by blast
+ then show ?case
+ proof
+ assume "x=x2 \<and> y=y2"
+ then show ?case using 2 by simp
+ next
+ assume ne:"\<not>(x=x2 \<and> y=y2)"
+ then show ?case
+ proof -
+ have "x\<le>x2 \<or> x>x2" by force
+ then show ?case
+ proof
+ assume xs: "x\<le>x2"
+ have "y\<le>y2 \<or> y>y2" by force
+ then show ?case
+ proof
+ assume ys: "y\<le>y2"
+ have swi:"invar sw" using 2 by simp
+ have nwi:"invar nw" using 2 by simp
+ have nei:"invar ne" using 2 by simp
+ have sei: "invar se" using 2 by simp
+
+ have "invar (qinsert (x, y) v sw)" using 2(4) xs ys ne swi by auto
+
+ have swo: "swOf sw (x2,y2)" using 2(5) by simp
+
+ have ksq: "(\<forall>p \<in> (keyset_quadtree sw) . issw p (x2,y2))" apply (cases sw) apply simp using swo 2(5) by simp
+
+ let ?qi = "(qinsert (x, y) v sw)"
+ have a: "swOf ?qi (x2,y2) = (\<forall>p \<in> (keyset_quadtree ?qi) . issw p (x2,y2)) " using qinsert_non_Leaf[of x y v sw] by (smt swOf.elims(2) swOf.elims(3))
+ also have b:"... = (\<forall>p \<in> (keyset_quadtree sw \<union> {(x,y)}) . issw p (x2,y2))" using ks[of x y v sw] by simp
+ also have c: "... = True" using ksq xs ys ne by simp
+
+ have d:"invar (qinsert (x, y) v (Node (x2, y2) v2 nw ne se sw)) = invar ( (Node (x2, y2) v2 nw ne se (qinsert (x, y) v sw)))" using xs ys ne by auto
+ have e: "... = True " using a b c 2 by force
+
+ from d e show ?case by simp
+ next
+ assume ys: "y2 < y"
+ have swi:"invar sw" using 2 by simp
+ have nwi:"invar nw" using 2 by simp
+ have nei:"invar ne" using 2 by simp
+ have sei: "invar se" using 2 by simp
+
+ have "invar (qinsert (x, y) v nw)" using 2 xs ys swi by auto
+
+ have nwo: "nwOf nw (x2,y2)" using 2 by simp
+
+ have ksq: "(\<forall>p \<in> (keyset_quadtree nw) . isnw p (x2,y2))" apply (cases nw) apply simp using nwo 2 by simp
+
+ let ?qi = "(qinsert (x, y) v nw)"
+ have a: "nwOf ?qi (x2,y2) = (\<forall>p \<in> (keyset_quadtree ?qi) . isnw p (x2,y2)) " using qinsert_non_Leaf[of x y v nw] by (smt nwOf.elims(2) nwOf.elims(3))
+ also have b:"... = (\<forall>p \<in> (keyset_quadtree nw \<union> {(x,y)}) . isnw p (x2,y2))" using ks[ of x y v nw] by simp
+ also have c: "... = True" using ksq xs ys by simp
+
+ have d:"invar (qinsert (x, y) v (Node (x2, y2) v2 nw ne se sw)) = invar ( (Node (x2, y2) v2 (qinsert (x, y) v nw) ne se sw))" using xs ys by simp
+ have e: "... = True " using a b c 2 by force
+ from d e show ?case by simp
+ qed
+ next
+ assume xs: "x>x2"
+ have "y\<le>y2 \<or> y>y2" by force
+ then show ?case
+ proof
+ assume ys: "y\<le>y2"
+ have swi:"invar sw" using 2 by simp
+ have nwi:"invar nw" using 2 by simp
+ have nei:"invar ne" using 2 by simp
+ have sei: "invar se" using 2 by simp
+
+ have "invar (qinsert (x, y) v se)" using 2(3) xs ys sei by auto
+
+ have seo: "seOf se (x2,y2)" using 2(5) by simp
+
+ have ksq: "(\<forall>p \<in> (keyset_quadtree se) . isse p (x2,y2))" apply (cases se) apply simp using seo 2(5) by simp
+
+ let ?qi = "(qinsert (x, y) v se)"
+ have a: "seOf ?qi (x2,y2) = (\<forall>p \<in> (keyset_quadtree ?qi) . isse p (x2,y2)) " using qinsert_non_Leaf[of x y v se] by (smt seOf.elims(2) seOf.elims(3))
+ also have b:"... = (\<forall>p \<in> (keyset_quadtree se \<union> {(x,y)}) . isse p (x2,y2))" using ks[ of x y v se] by simp
+ also have c: "... = True" using ksq xs ys by simp
+
+ have d:"invar (qinsert (x, y) v (Node (x2, y2) v2 nw ne se sw)) = invar ( (Node (x2, y2) v2 nw ne (qinsert (x, y) v se) sw))" using xs ys by simp
+ have e: "... = True " using a b c 2 by force
+
+ from d e show ?case by simp
+ next
+ assume ys: "y2 < y"
+ have swi:"invar sw" using 2 by simp
+ have nwi:"invar nw" using 2 by simp
+ have nei:"invar ne" using 2 by simp
+ have sei: "invar se" using 2 by simp
+
+ have "invar (qinsert (x, y) v ne)" using 2 xs ys nei by auto
+
+ have neo: "neOf ne (x2,y2)" using 2 by simp
+
+ have ksq: "(\<forall>p \<in> (keyset_quadtree ne) . isne p (x2,y2))" apply (cases ne) apply simp using neo 2 by simp
+
+ let ?qi = "(qinsert (x, y) v ne)"
+ have a: "neOf ?qi (x2,y2) = (\<forall>p \<in> (keyset_quadtree ?qi) . isne p (x2,y2)) " using qinsert_non_Leaf[of x y v ne] by (smt neOf.elims(2) neOf.elims(3))
+ also have b:"... = (\<forall>p \<in> (keyset_quadtree ne \<union> {(x,y)}) . isne p (x2,y2))" using ks[ of x y v ne] by simp
+ also have c: "... = True" using ksq xs ys by simp
+
+ have d:"invar (qinsert (x, y) v (Node (x2, y2) v2 nw ne se sw)) = invar ( (Node (x2, y2) v2 nw (qinsert (x, y) v ne) se sw))" using xs ys by simp
+ have e: "... = True " using a b c 2 by force
+ from d e show ?case by simp
+ qed
+ qed
+ qed
+ qed
+qed
+
+lemma ins_set: "(x,y)\<notin>keyset_quadtree t \<Longrightarrow> set_quadtree(qinsert (x,y) v t) = set_quadtree(t) \<union> {v}"
+ apply (induction t)
+ apply force
+ apply force
+ done
+
+lemma "invar t\<Longrightarrow> mapset_quadtree(qinsert (x,y) v t) = {((xe,ye),v)\<in> mapset_quadtree(t). \<not>(xe=x \<and>ye=y)} \<union> {((x,y), v)}"
+ sorry
+
+(*Tests*)
+value "qinsert (0,0) 1 (qinsert (2,2) 3 (qinsert (1,3) 3 Leaf))"
+value "list_quadtree (qinsert (0,0) 1 (qinsert (2,2) 3 (qinsert (1,3) 3 Leaf)))"
+value "delete (1,3) (qinsert (0,0) 1 (qinsert (2,2) 3 (qinsert (1,3) 3 Leaf)))"
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Erhard_Julian_ga48kog@mytum.de_655/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,2 @@
+Using temporary directory '/tmp/tmp.8G3jCZptOH'
+Exercise path '/home/admin/suites/exbeoriginal' not found
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Groer_Markus_markus.grosser@tum.de_636/Arbitrary_Precision.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,236 @@
+theory Arbitrary_Precision
+ imports Main
+begin
+
+text \<open>
+ Arbitrary precision arithmetic on lists of bits. Only multiplication and addition have been
+ implemented and proven.
+\<close>
+
+section \<open>Type Definition and Conversions\<close>
+
+type_synonym bitlist = "bool list"
+
+fun bitlist_of_nat :: "nat \<Rightarrow> bitlist" where
+ "bitlist_of_nat 0 = []"
+| "bitlist_of_nat n = odd n # bitlist_of_nat (n div 2)"
+
+fun nat_of_bitlist :: "bitlist \<Rightarrow> nat" where
+ "nat_of_bitlist [] = 0"
+| "nat_of_bitlist (b#bs) = (if b then 1 else 0) + 2 * nat_of_bitlist bs"
+
+section \<open>Optimality\<close>
+
+text \<open>
+ A bitlist is considered optimal if it has no leading zeroes.
+\<close>
+
+definition optimal :: "bitlist \<Rightarrow> bool" where
+ "optimal bl = (last bl \<or> bl = [])"
+
+lemma optimal_cons[simp]: "bs \<noteq> [] \<Longrightarrow> optimal (b#bs) = optimal bs"
+ unfolding optimal_def by simp
+
+lemma optimal_nat_of_bitlist_ne_zero: "\<lbrakk>optimal bs; bs \<noteq> []\<rbrakk> \<Longrightarrow> nat_of_bitlist bs \<noteq> 0"
+ by (induction bs; fastforce simp: optimal_def)
+
+lemma nat_of_bitlist_bitlist_of_nat[simp]: "nat_of_bitlist (bitlist_of_nat n) = n"
+proof (induction n rule: less_induct)
+ case (less x)
+ then show ?case by (cases x) auto
+qed
+
+lemma bitlist_of_nat_nat_of_bitlist[simp]: "optimal bl \<Longrightarrow> bitlist_of_nat (nat_of_bitlist bl) = bl"
+proof (induction bl)
+ case Nil
+ then show ?case by simp
+next
+ case (Cons b bs)
+ then show ?case proof (cases b)
+ case True
+ then have "optimal bs" using Cons unfolding optimal_def by auto
+ then show ?thesis using Cons True by auto
+ next
+ case False
+ then have "bs \<noteq> []" using Cons.prems unfolding optimal_def by auto
+ then have "nat_of_bitlist bs \<noteq> 0"
+ using Cons.prems by (simp add: optimal_nat_of_bitlist_ne_zero)
+ then have "2 * nat_of_bitlist bs \<noteq> 0" by simp
+ then have "\<exists>n. 2 * nat_of_bitlist bs = Suc n" by presburger
+ moreover have "\<not> odd (2 * nat_of_bitlist bs)" by simp
+ moreover have "(2 * nat_of_bitlist bs) div 2 = nat_of_bitlist bs" by simp
+ ultimately have
+ "bitlist_of_nat (2 * nat_of_bitlist bs) = False # bitlist_of_nat (nat_of_bitlist bs)"
+ by force
+ then show ?thesis using Cons False by (auto simp: \<open>bs \<noteq> []\<close>)
+ qed
+qed
+
+section \<open>Addition\<close>
+
+text \<open>
+ A full adder, as is commonly implemented in digital logic:
+\<close>
+
+abbreviation full_adder :: "bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool * bool" where
+ "full_adder a b c \<equiv> (let ab = a \<noteq> b in ((ab \<noteq> c), (a \<and> b) \<or> (ab \<and> c)))"
+
+fun add_with_carry :: "bitlist \<Rightarrow> bitlist \<Rightarrow> bool \<Rightarrow> bitlist" where
+ "add_with_carry (a#as) (b#bs) c = (let (r, c') = full_adder a b c in r # add_with_carry as bs c')"
+| "add_with_carry (a#as) [] c = (a \<noteq> c) # add_with_carry as [] (a \<and> c)"
+| "add_with_carry [] (b#bs) c = (b \<noteq> c) # add_with_carry [] bs (b \<and> c)"
+| "add_with_carry [] [] c = (if c then [True] else [])"
+
+lemma add_with_carry_empty: "add_with_carry x y c = [] \<Longrightarrow> x = [] \<and> y = [] \<and> \<not>c"
+ by (induction x y c rule: add_with_carry.induct)
+ (auto simp: Let_def)
+
+lemma optimal_add_with_carry[simp]: "\<lbrakk>optimal x; optimal y\<rbrakk> \<Longrightarrow> optimal (add_with_carry x y c)"
+proof (induction x y c rule: add_with_carry.induct)
+ case (1 a as b bs c)
+ let ?r = "fst (full_adder a b c)"
+ let ?c' = "snd (full_adder a b c)"
+ show ?case proof (cases ?c')
+ case True
+ then have "add_with_carry as bs ?c' \<noteq> []" using add_with_carry_empty by meson
+ then show ?thesis
+ using 1 by (auto simp: optimal_def split: if_splits prod.splits)
+ next
+ case False
+ then show ?thesis
+ using 1 optimal_def apply (auto split: prod.splits)
+ using add_with_carry_empty(2)[of as bs False] by fastforce+
+ qed
+next
+ case (2 a as c)
+ then have "\<lbrakk>add_with_carry as [] (a \<and> c) = []; as = []\<rbrakk> \<Longrightarrow> optimal (add_with_carry (a # as) [] c)"
+ using optimal_def by force
+ then show ?case
+ using 2 neq_Nil_conv[of as] neq_Nil_conv[of "add_with_carry as [] (a \<and> c)"] optimal_cons
+ by auto
+next
+ case (3 b bs c)
+ then have "\<lbrakk>add_with_carry [] bs (b \<and> c) = []; bs = []\<rbrakk> \<Longrightarrow> optimal (add_with_carry [] (b # bs) c)"
+ using optimal_def by force
+ then show ?case
+ using 3 neq_Nil_conv[of bs] neq_Nil_conv[of "add_with_carry [] bs (b \<and> c)"] optimal_cons
+ by auto
+next
+ case (4 c)
+ then show ?case by (simp add: optimal_def)
+qed
+
+lemma nat_of_bitlist_add_with_carry[simp]: "nat_of_bitlist (add_with_carry x y c) = nat_of_bitlist x + nat_of_bitlist y + (if c then 1 else 0)"
+ by (induction x y c rule: add_with_carry.induct) auto
+
+definition add_bitlist :: "bitlist \<Rightarrow> bitlist \<Rightarrow> bitlist" where
+ "add_bitlist x y = add_with_carry x y False"
+
+lemma optimal_add_bitlist[simp]: "\<lbrakk>optimal x; optimal y\<rbrakk> \<Longrightarrow> optimal (add_bitlist x y)"
+ unfolding add_bitlist_def by simp
+
+lemma nat_of_bitlist_add_bitlist[simp]: "nat_of_bitlist (add_bitlist x y) = nat_of_bitlist x + nat_of_bitlist y"
+ unfolding add_bitlist_def by simp
+
+section \<open>Multiplication\<close>
+
+text \<open>
+ Naive \(\mathcal{O}(n²)\) implementation, with \(n\) being the number of bits. Since lists don't
+ allow fast random access, I am not even sure how much better Karatsuba or FFT multiplication would
+ fare.
+\<close>
+
+fun mult_bitlist :: "bitlist \<Rightarrow> bitlist \<Rightarrow> bitlist" where
+ "mult_bitlist (a#as) bs =
+ (let r = case mult_bitlist as bs of [] \<Rightarrow> [] | r \<Rightarrow> False # r in
+ if a then add_bitlist bs r else r)"
+| "mult_bitlist [] _ = []"
+
+lemma optimal_mult_bitlist[simp]: "\<lbrakk>optimal x; optimal y\<rbrakk> \<Longrightarrow> optimal (mult_bitlist x y)"
+proof (induction x)
+ case (Cons a as)
+ then have "optimal as" by (cases as) (auto simp: optimal_def)
+ then show ?case using Cons by (auto simp: Let_def split: list.splits)
+qed auto
+
+lemma nat_of_bitlist_mult_bitlist[simp]: "nat_of_bitlist (mult_bitlist x y) = nat_of_bitlist x * nat_of_bitlist y"
+ by (induction x) (auto split: list.splits)
+
+section \<open>Bitshifts\<close>
+
+fun shl_bitlist :: "bitlist \<Rightarrow> nat \<Rightarrow> bitlist" where
+ "shl_bitlist [] _ = []"
+| "shl_bitlist bl n = replicate n False @ bl"
+
+lemma optimal_shl_bitlist[simp]: "optimal x \<Longrightarrow> optimal (shl_bitlist x n)"
+ unfolding optimal_def by (cases x) auto
+
+lemma nat_of_bitlist_shl_bitlist[simp]: "nat_of_bitlist (shl_bitlist x n) = 2 ^ n * nat_of_bitlist x"
+ by (induction n; cases x) auto
+
+definition shr_bitlist :: "bitlist \<Rightarrow> nat \<Rightarrow> bitlist" where
+ "shr_bitlist bl n = drop n bl"
+
+lemma optimal_shr_bitlist[simp]: "optimal x \<Longrightarrow> optimal (shr_bitlist x n)"
+ unfolding shr_bitlist_def optimal_def by (cases x) auto
+
+lemma nat_of_bitlist_shr_bitlist[simp]: "nat_of_bitlist (shr_bitlist x n) = nat_of_bitlist x div 2 ^ n"
+ unfolding shr_bitlist_def proof (induction n arbitrary: x)
+ case (Suc n)
+ then show ?case by (cases x) (auto simp: algebra_simps Divides.div_mult2_eq)
+qed auto
+
+section \<open>Comparisons\<close>
+
+text \<open>
+ I am pretty sure this type is already defined somewhere:
+\<close>
+
+datatype ordering = EQ | LT | GT
+
+fun cmp_bitlist :: "bitlist \<Rightarrow> bitlist \<Rightarrow> ordering" where
+ "cmp_bitlist [] [] = EQ"
+| "cmp_bitlist [] (_#_) = LT"
+| "cmp_bitlist (_#_) [] = GT"
+| "cmp_bitlist (a#as) (b#bs) =
+ (case cmp_bitlist as bs of
+ EQ \<Rightarrow> if a = b then EQ else if a then GT else LT
+ | r \<Rightarrow> r)"
+
+lemma cmp_bitlist_eq: "\<lbrakk>optimal x; optimal y\<rbrakk> \<Longrightarrow> cmp_bitlist x y = EQ \<longleftrightarrow> nat_of_bitlist x = nat_of_bitlist y"
+proof (induction x y rule: cmp_bitlist.induct)
+ case (4 a as b bs)
+ then show ?case by (auto split: ordering.splits)
+ (fastforce simp: optimal_def double_not_eq_Suc_double Suc_double_not_eq_double)+
+qed (fastforce dest: optimal_nat_of_bitlist_ne_zero)+
+
+lemma cmp_bitlist_lt: "\<lbrakk>optimal x; optimal y\<rbrakk> \<Longrightarrow> cmp_bitlist x y = LT \<longleftrightarrow> nat_of_bitlist x < nat_of_bitlist y"
+proof (induction x y rule: cmp_bitlist.induct)
+ case (4 a as b bs)
+ then show ?case
+ apply (cases a; cases b, auto split: ordering.splits)
+ proof goal_cases
+ case 7 then show ?case by (metis cmp_bitlist_eq lessI optimal_cons optimal_def)
+ next
+ case 9 then show ?case
+ by (metis Suc_1 cmp_bitlist_eq linorder_neqE_nat mult_strict_left_mono not_less_eq
+ optimal_cons optimal_def ordering.distinct(3) zero_less_Suc)
+ qed (fastforce simp: optimal_def)+ (* Dear god I hate myself for this proof *)
+qed (fastforce dest: optimal_nat_of_bitlist_ne_zero)+
+
+lemma cmp_bitlist_gt: "\<lbrakk>optimal x; optimal y\<rbrakk> \<Longrightarrow> cmp_bitlist x y = GT \<longleftrightarrow> nat_of_bitlist x > nat_of_bitlist y"
+proof (induction x y rule: cmp_bitlist.induct)
+ case (4 a as b bs)
+ then show ?case proof (auto split: ordering.splits, goal_cases)
+ case 2
+ then show ?case by (metis cmp_bitlist_eq lessI optimal_cons optimal_def)
+ next
+ case 6
+ then show ?case
+ by (metis cmp_bitlist_eq linorder_neqE_nat mult_strict_left_mono not_less_eq optimal_cons
+ optimal_def ordering.distinct(1) zero_less_numeral)
+ qed (fastforce simp: optimal_def)+ (* Same shit as above :/ *)
+qed (fastforce dest: optimal_nat_of_bitlist_ne_zero)+
+
+end
+
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Groer_Markus_markus.grosser@tum.de_636/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,2 @@
+Using temporary directory '/tmp/tmp.ALSOLpXENs'
+Exercise path '/home/admin/suites/exbeoriginal' not found
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Haslbeck_Maximilian_max.haslbeck@mytum.de_640/Pell.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,933 @@
+(*\<Rightarrow>Manuel*)
+theory Pell
+ imports Analysis
+begin
+
+(*
+The equation
+x^2 + n * y^2 = 1
+has integer solutions in x and y for every non-square n.
+*)
+
+(*
+- Only works with isabelle-dev
+- should also work with Complex Main, but first needs some minor refinements
+*)
+
+(* Lemma by Manuel Eberl *)
+lemma nonneg_root_nat_or_irrat:
+ assumes "x ^ 2 = real a" and "x \<ge> 0"
+ shows "x \<in> \<nat> \<or> x \<notin> \<rat>"
+proof safe
+ assume "x \<notin> \<nat>" and "x \<in> \<rat>"
+ from Rats_abs_nat_div_natE[OF this(2)]
+ obtain p q :: nat where q_nz [simp]: "q \<noteq> 0" and "abs x = p / q" and coprime: "coprime p q" .
+ with \<open>x \<ge> 0\<close> have x: "x = p / q"
+ by simp
+ with assms have "real (q ^ 2) * real a = real (p ^ 2)"
+ by (simp add: field_simps)
+ also have "real (q ^ 2) * real a = real (q ^ 2 * a)"
+ by simp
+ finally have "p ^ 2 = q ^ 2 * a"
+ by (subst (asm) of_nat_eq_iff) auto
+ hence "q ^ 2 dvd p ^ 2"
+ by simp
+ hence "q dvd p"
+ by (rule pow_divs_pow) auto
+ with coprime have "q = 1"
+ by simp
+ with x and \<open>x \<notin> \<nat>\<close> show False
+ by simp
+qed
+
+(* Lemma by Manuel Eberl *)
+lemma minus_in_Ints_iff [simp]:
+ "-x \<in> \<int> \<longleftrightarrow> x \<in> \<int>"
+ using Ints_minus[of x] Ints_minus[of "-x"] by auto
+
+(* Lemma by Manuel Eberl *)
+lemma root_nat_or_irrat:
+ assumes "x ^ 2 = real a"
+ shows "x \<in> \<int> \<or> x \<notin> \<rat>"
+proof (cases "x \<ge> 0")
+ case True
+ with nonneg_root_nat_or_irrat[OF assms this]
+ show ?thesis by (auto simp: Nats_altdef2)
+next
+ case False
+ from assms have "(-x) ^ 2 = real a"
+ by simp
+ moreover from False have "-x \<ge> 0"
+ by simp
+ ultimately have "-x \<in> \<nat> \<or> -x \<notin> \<rat>"
+ by (rule nonneg_root_nat_or_irrat)
+ thus ?thesis
+ by (auto simp: Nats_altdef2)
+qed
+
+lemma sqrt_nat_or_irrat:
+ "sqrt (real a) \<in> \<nat> \<or> sqrt (real a) \<notin> \<rat>"
+ using nonneg_root_nat_or_irrat[of "sqrt a" a] by auto
+
+lemma not_Rats_mult: fixes x y :: real
+ assumes "x \<noteq> 0" "x \<in> \<rat>" "y \<notin> \<rat>"
+ shows "x * y \<notin> \<rat>"
+proof (rule ccontr)
+ assume 0: "\<not> x * y \<notin> \<rat>"
+ have "(x * y) / x \<in> \<rat>"
+ apply(rule Rats_divide)
+ using 0 assms by auto
+ also have "(x * y) / x = y"
+ using assms by(simp)
+ ultimately show False using assms by simp
+qed
+
+
+lemma not_Rats_add: fixes x y :: real
+ assumes "x \<in> \<rat>" "y \<notin> \<rat>"
+ shows "x + y \<notin> \<rat>" "y + x \<notin> \<rat>"
+ using assms Rats_diff by force+
+
+lemma not_Rats_diff: fixes x y :: real
+ assumes "x \<in> \<rat>" "y \<notin> \<rat>"
+ shows "x - y \<notin> \<rat>" "y - x \<notin> \<rat>"
+ using Rats_diff assms Rats_add by force+
+
+lemma not_Rats_frac: fixes x :: real
+ shows "x \<in> \<rat> \<longleftrightarrow> frac x \<in> \<rat>"
+ apply(rule)
+ unfolding frac_def
+ apply simp
+ using Rats_of_int not_Rats_diff by blast
+
+lemma pigeonhole:
+ assumes "finite B" "card A > card B" "A \<subseteq> \<Union>B"
+ shows "\<exists>x y M. x \<noteq> y \<and> x \<in> A \<and> y \<in> A \<and> M \<in> B \<and> x \<in> M \<and> y \<in> M"
+proof(rule ccontr)
+ assume "\<nexists>x y M. x \<noteq> y \<and> x \<in> A \<and> y \<in> A \<and> M \<in> B \<and> x \<in> M \<and> y \<in> M"
+ then have "\<And>x y M. x \<noteq> y \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> M \<in> B \<Longrightarrow> x \<in> M \<Longrightarrow> y \<notin> M"
+ by blast
+ with assms show False
+ proof(induction "card B" arbitrary: A B)
+ case (Suc x)
+ then have "\<exists>b. b \<in> B"
+ by (metis Suc_n_not_le_n card.empty empty_iff subsetI subset_antisym zero_le)
+ then obtain b where "b \<in> B"
+ by blast
+ with Suc have 0: "card (B - {b}) = card B - 1"
+ by force
+ show ?case
+ proof(cases "A \<inter> b = {}")
+ case True
+ show ?thesis
+ apply(rule Suc(1)[of "B - {b}" A])
+ using 0 Suc apply(simp)
+ using Suc apply(simp)
+ apply(subst 0) using Suc apply(simp)
+ using Suc True apply(fast)
+ using Suc
+ by blast
+ next
+ case False
+ then have "\<exists>a. a \<in> A \<and> a \<in> b"
+ by auto
+ then obtain a where "a \<in> A" "a \<in> b"
+ by blast
+ then have 3: "\<And>a'. a' \<in> A - {a} \<Longrightarrow> a' \<notin> b"
+ using Suc `b \<in> B` by blast
+ show ?thesis
+ apply(rule Suc(1)[of "B - {b}" "A - {a}"])
+ using "0" Suc.hyps(2) apply linarith
+ apply (simp add: Suc.prems(1))
+ apply (metis "0" Suc.hyps(2) Suc.prems(2) Suc_lessD Suc_less_SucD \<open>a \<in> A\<close> card_Suc_Diff1 card_infinite diff_Suc_1 finite.emptyI finite.insertI finite_Diff2)
+ defer
+ using Suc.prems(4) apply auto[1]
+ using Suc(5) `a \<in> b` 3
+ by blast
+ qed
+ qed (auto)
+qed
+
+lemma split_that: fixes x y D p q::"'a::field"
+ shows "(x^2 - D * y^2) * (p^2 - D * q^2) = (x*p + D*y*q)^2 - D * (x*q + y*p)^2"
+ by (auto simp add: algebra_simps power2_sum power_mult_distrib semiring_normalization_rules(29))
+
+locale pell =
+ fixes d::real
+ assumes d_greater_zero: "0 < d" and
+ d_non_square: "\<nexists>x. x \<in> \<int> \<and> x ^ 2 = d" and
+ d_nat: "d \<in> \<nat>"
+begin
+
+lemma sqrt_d_not_int: "sqrt d \<notin> \<int>"
+ using d_non_square
+ using d_greater_zero less_eq_real_def real_sqrt_pow2 by blast
+
+lemma sqrt_d_not_rat: "sqrt d \<notin> \<rat>"
+ by (metis Ints_of_nat Nats_cases d_nat sqrt_d_not_int sqrt_nat_or_irrat)
+
+lemma d_greater_eq_2: "d \<ge> 2"
+proof -
+ have "sqrt d \<noteq> 1"
+ using sqrt_d_not_int by force
+ then have "d \<noteq> 1"
+ by simp
+ show "d \<ge> 2"
+ apply(rule Nats_cases[OF d_nat])
+ using d_greater_zero `d \<noteq> 1` by auto
+qed
+
+lemma sqrt_d_greater_1: "sqrt d > 1"
+ using real_sqrt_gt_1_iff d_greater_eq_2 by simp
+
+lemma how_should_i_name_this_one:
+ assumes "(q::nat) > s" "(x::real) \<ge> 1"
+ shows "0 < \<lfloor>q * x\<rfloor> - \<lfloor>s * x\<rfloor>"
+proof -
+ have 0: "\<lfloor>s * x\<rfloor> \<le> s * x"
+ by simp
+ have "0 \<le> x - 1"
+ using assms by simp
+ also have "\<dots> \<le> (int q - int s) * x - 1"
+ using assms by fastforce
+ also have "\<dots> \<le> q * x - 1 - s * x"
+ by (simp add: left_diff_distrib')
+ also have "\<dots> \<le> q * x - 1 - \<lfloor>s * x\<rfloor>"
+ by force
+ also have "\<dots> < \<lfloor>q * x\<rfloor> - \<lfloor>s * x\<rfloor>"
+ by simp
+ finally show ?thesis
+ by simp
+qed
+
+lemma floor_int: "x \<in> \<int> \<longleftrightarrow> floor x = x"
+ by (metis (no_types, hide_lams) Ints_of_int add.inverse_inverse add_uminus_conv_diff
+cancel_comm_monoid_add_class.diff_cancel floor_uminus_of_int
+less_add_same_cancel1 not_le not_less_iff_gr_or_eq of_int_minus of_int_of_nat_eq of_nat_nat)
+
+lemma "x \<notin> \<int> \<longleftrightarrow> floor x < x"
+ using floor_int by force+
+
+lemma inj_real_of_int: "inj real_of_int"
+ apply(rule)
+ by simp
+
+
+lemma always_smaller:
+ fixes N::nat
+ assumes "N > 0"
+ shows "\<exists>t::int. \<exists>u::int. t > 0 \<and> u > 0 \<and> \<bar>t - u * sqrt d\<bar> < 1 / N \<and> 1 / N \<le> 1/\<bar>u\<bar>"
+proof -
+ define interval_N where "interval_N = (\<lambda>i. { i / N <..< (i + 1) / N})"
+ define intervals_0_1 where "intervals_0_1 = interval_N ` {0..N - 1}"
+ have 0: "{0..1} - (\<lambda>x. x / N) ` {0..N} \<subseteq> \<Union> intervals_0_1"
+ unfolding intervals_0_1_def interval_N_def apply(clarsimp)
+ proof (goal_cases)
+ case (1 x)
+ note omei = 1
+ have x1: "x < 1"
+ proof(rule ccontr)
+ assume "\<not> x < 1"
+ with 1 have "x = 1"
+ by simp
+ then have "x = N / N"
+ using assms by simp
+ then have "x \<in> (\<lambda>x. real x / real N) ` {0..N}"
+ by fastforce
+ then show False
+ using 1 by simp
+ qed
+ have x0: "0 < x"
+ proof(rule ccontr)
+ assume "\<not> 0 < x"
+ with 1 have "x = 0"
+ by simp
+ then have "x = 0 / N"
+ using assms by simp
+ then have "x \<in> (\<lambda>x. real x / real N) ` {0..N}"
+ by fastforce
+ then show False
+ using 1 by simp
+ qed
+ from x1 x0 have f: "floor (x * N) \<in> {0..N - Suc 0}"
+ apply(auto)
+ by (metis Groups.add_ac(2) Suc_pred assms le_floor_iff mult.left_neutral not_less
+ of_int_of_nat_eq of_nat_0_less_iff real_mult_less_iff1 semiring_1_class.of_nat_simps(2) zle_add1_eq_le)
+ have "floor (x * N) \<le> x * N"
+ by simp
+ moreover have "floor (x * N) \<noteq> x * N"
+ proof(rule ccontr, goal_cases)
+ case 1
+ then have "\<exists>n. x * real N = real n"
+ by (metis \<open>\<lfloor>x * real N\<rfloor> \<in> {0..int (N - Suc 0)}\<close> atLeastAtMost_iff nonneg_int_cases of_int_of_nat_eq)
+ then obtain n where n_def: "x * real N = real n"
+ by blast
+ have "n \<le> N"
+ by (metis \<open>x * real N = real n\<close> linordered_semidom_class.of_nat_le_iff mult_left_le_one_le of_nat_0_le_iff omei(2) omei(3))
+ have "x = n / N"
+ using n_def assms
+ by (simp add: nonzero_eq_divide_eq)
+ then have "x \<in> (\<lambda>x. real x / real N) ` {0..N}"
+ using `n \<le> N` n_def by force
+ then show ?case using omei by blast
+ qed
+ ultimately have "floor (x * N) < x * N"
+ by linarith
+ then have "(floor (x * N)) / N < x"
+ using assms
+ by (simp add: divide_less_eq)
+ have "x * N < floor (x * N) + 1"
+ by simp
+ then have "x < (floor (x * N) + 1) / N"
+ using assms
+ by (metis nonzero_eq_divide_eq of_nat_0_less_iff order_less_irrefl pos_divide_less_eq)
+ have "floor (x * N) \<in> \<nat>"
+ unfolding Nats_altdef2 apply(auto)
+ apply (simp add: floor_int)
+ by (simp add: omei(2))
+ then have "\<exists>n. floor (x * N) = real n"
+ by (metis Nats_cases of_int_of_nat_eq)
+ then obtain n where n_def: "real n = floor (x * N)"
+ by fastforce
+ then have 3: "x < (real n + 1) / N"
+ using `x < (floor (x * N) + 1) / N` by (simp add: add.commute)
+ from n_def have 4: "n \<in> {0..N - 1}"
+ using f by fastforce
+ have 5: "n / N < x"
+ using n_def `(floor (x * N)) / N < x` by simp
+ show ?case using 3 4 5
+ by auto
+ qed
+ have 1: "(\<lambda>x. frac (x * sqrt d)) ` {1..N+1} \<subseteq> \<Union> intervals_0_1"
+ apply(rule subset_trans[OF _ 0]) apply(rule) apply(rule)
+ apply(auto)[1]
+ apply (simp add: frac_lt_1 less_eq_real_def)
+ subgoal for x
+ apply(subgoal_tac "x \<notin> \<rat>")
+ subgoal apply(rule)
+ proof (goal_cases)
+ case 1
+ from 1(3) have "\<exists>n \<in> {0..N}. n / N = x"
+ by blast
+ show ?case
+ using "1"(2) Rats_divide Rats_of_nat \<open>\<exists>n\<in>{0..N}. real n / real N = x\<close> by blast
+ qed
+ apply(auto)
+ apply(subst (asm) not_Rats_frac[symmetric])
+ by (simp add: not_Rats_mult sqrt_d_not_rat)
+ done
+ have "card intervals_0_1 = N"
+ unfolding intervals_0_1_def interval_N_def
+ apply(subst card_image)
+ subgoal apply(rule) proof(auto, goal_cases)
+ case (1 x y)
+ have "Inf {real x / real N<..<(real x + 1) / real N} = real x / real N"
+ using `N > 0` by (auto intro!: cInf_greaterThanLessThan divide_strict_right_mono)
+ moreover have "Inf {real y / real N<..<(real y + 1) / real N} = real y / real N"
+ using `N > 0` by (auto intro!: cInf_greaterThanLessThan divide_strict_right_mono)
+ ultimately have "real x / real N = real y / real N"
+ using 1 by auto
+ then show ?case
+ using `N > 0` 1 by (auto)
+ qed
+ using `N > 0` by (induction N) auto
+ also have "card ((\<lambda>x. frac (x * sqrt d)) ` {1..N+1}) = N + 1"
+ apply(subst card_image)
+ subgoal apply(rule)
+ proof (rule ccontr, goal_cases)
+ case (1 x y)
+ then have "frac (real x * sqrt d) - frac (real y * sqrt d) = 0"
+ by fastforce
+ then have "(real x - real y) * sqrt d - \<lfloor>real x * sqrt d\<rfloor> + \<lfloor>real y * sqrt d\<rfloor> = 0"
+ unfolding frac_def by algebra
+ also have "(real x - real y) * sqrt d - \<lfloor>real x * sqrt d\<rfloor> + \<lfloor>real y * sqrt d\<rfloor> \<notin> \<rat>"
+ apply(subst add.commute)
+ apply(rule not_Rats_add)
+ apply(simp)
+ apply(rule not_Rats_diff(2))
+ apply(simp)
+ apply(rule not_Rats_mult)
+ using 1 apply(simp)
+ apply(rule Rats_diff)
+ using 1 Rats_of_nat apply blast
+ using 1 Rats_of_nat apply(blast)
+ by (simp add: sqrt_d_not_rat)
+ ultimately have False
+ by fastforce
+ then show ?case ..
+ qed
+ by simp
+ ultimately have 2: "card ((\<lambda>x. frac (x * sqrt d)) ` {1..N+1}) > card intervals_0_1"
+ by simp
+ have "\<exists>x y M.
+ x \<noteq> y \<and>
+ x \<in> (\<lambda>x. frac (real x * sqrt d)) ` {1..N + 1} \<and>
+ y \<in> (\<lambda>x. frac (real x * sqrt d)) ` {1..N + 1} \<and>
+ M \<in> intervals_0_1 \<and> x \<in> M \<and> y \<in> M"
+ apply(rule pigeonhole)
+ apply(simp add: intervals_0_1_def)
+ using 1 2 by(auto)
+ then obtain m1 m2 M where m1m2: "m1 \<noteq> m2"
+ "m1 \<in> (\<lambda>x. frac (real x * sqrt d)) ` {1..N + 1}"
+ "m2 \<in> (\<lambda>x. frac (real x * sqrt d)) ` {1..N + 1}"
+ and
+ M: "M \<in> intervals_0_1"
+ "m1 \<in> M"
+ "m2 \<in> M"
+ by blast
+ from M have "\<exists>n. n \<in> {0..N - 1} \<and> interval_N (real n) = M"
+ unfolding intervals_0_1_def by blast
+ then obtain n where n_def: "n \<in> {0..N - 1}" "M = interval_N (real n)"
+ by blast
+ have "\<And>x y z. x \<in> interval_N z \<Longrightarrow> y \<in> interval_N z \<Longrightarrow> \<bar>x - y\<bar> < \<bar>(z + 1) / N - z / N\<bar>"
+ unfolding interval_N_def by auto
+ then have 431: "\<And>x y z. x \<in> interval_N z \<Longrightarrow> y \<in> interval_N z \<Longrightarrow> \<bar>x - y\<bar> < \<bar>1 / N\<bar>"
+ by (metis add_diff_cancel_left' diff_divide_distrib)
+ have "\<exists>q. q \<in> {1..N + 1} \<and> m1 = frac (q * sqrt d)"
+ "\<exists>s. s \<in> {1..N + 1} \<and> m2 = frac (s * sqrt d)"
+ using m1m2 by fast+
+ then obtain q s where qs: "q \<in> {1..N + 1} \<and> m1 = frac (q * sqrt d)"
+ "s \<in> {1..N + 1} \<and> m2 = frac (s * sqrt d)"
+ by blast+
+ have "q \<noteq> s"
+ using qs m1m2 by blast
+ have "\<bar> m1 - m2 \<bar> < \<bar>1 / N\<bar>"
+ apply(rule 431)
+ using n_def M by (auto)
+ then also have "\<bar> 1 / N \<bar> = 1 / N"
+ using assms by auto
+ ultimately have name_me: "\<bar> m1 - m2 \<bar> < 1 / N"
+ by simp
+ have "q \<noteq> s"
+ using qs m1m2 by blast
+ then show ?thesis
+ proof (rule linorder_neqE)
+ assume "q < s"
+ have 0: "\<bar> m2 - m1 \<bar> = \<bar>((floor (s * sqrt d)) - (floor (q * sqrt d))) - (int s - int q) * sqrt d\<bar>"
+ by (simp add: left_diff_distrib' frac_def qs)
+ define t where "t = ((floor (s * sqrt d)) - (floor (q * sqrt d)))"
+ define u where "u = (int s - int q)"
+ have 0: "\<bar>t - u * sqrt d \<bar> < 1 / N"
+ using name_me 0 t_def u_def by auto
+ have 1: "u > 0"
+ using `q < s` u_def by force
+ have 2: "t > 0"
+ unfolding t_def
+ apply(rule how_should_i_name_this_one)
+ using `q < s`
+ using sqrt_d_greater_1 by auto
+ have "s - q = int s - int q"
+ using `q < s` by force
+ also have "\<dots> \<le> N"
+ using qs by fastforce
+ finally have "u \<le> N"
+ using u_def by simp
+ then have 3: "1 / u \<ge> 1 / N"
+ by (simp add: "1" frac_le)
+ show ?thesis
+ using 0 1 2 3 by force
+ next
+ assume "s < q"
+ have 0: "\<bar> m1 - m2 \<bar> = \<bar>((floor (q * sqrt d)) - (floor (s * sqrt d))) - (int q - int s) * sqrt d\<bar>"
+ by (simp add: left_diff_distrib' frac_def qs)
+ define t where "t = ((floor (q * sqrt d)) - (floor (s * sqrt d)))"
+ define u where "u = (int q - int s)"
+ have 0: "\<bar>t - u * sqrt d \<bar> < 1 / N"
+ using name_me 0 t_def u_def by auto
+ have 1: "u > 0"
+ using `s < q` u_def by force
+ have 2: "t > 0"
+ unfolding t_def
+ apply(rule how_should_i_name_this_one)
+ using `s < q`
+ using sqrt_d_greater_1 by auto
+ have "q - s = int q - int s"
+ using `s < q` by force
+ also have "\<dots> \<le> N"
+ using qs by fastforce
+ finally have "u \<le> N"
+ using u_def by simp
+ then have 3: "1 / u \<ge> 1 / N"
+ by (simp add: "1" frac_le)
+ show ?thesis
+ using 0 1 2 3 by force
+ qed
+qed
+
+lemma pell_not_zero: fixes t u::int
+ assumes "u \<noteq> 0"
+ shows "t\<^sup>2 - d* u^2 \<noteq> 0"
+proof (rule ccontr)
+ assume a: "\<not> (t\<^sup>2) - d * (u\<^sup>2) \<noteq> 0"
+ then have "t \<noteq> 0"
+ using assms d_non_square by force
+ from a have "t^2 = d * u^2"
+ by simp
+ then have "sqrt (t\<^sup>2) = sqrt (d * (u\<^sup>2))"
+ by simp
+ then have "\<bar>t\<bar> = sqrt d * \<bar>u\<bar>"
+ using real_sqrt_mult by auto
+ moreover have "real_of_int \<bar>t\<bar> \<in> \<rat>"
+ using Rats_of_int by blast
+ moreover have "sqrt d * \<bar>u\<bar> \<notin> \<rat>"
+ apply(subst mult.commute)
+ apply(intro not_Rats_mult)
+ using `u \<noteq> 0` Rats_of_int sqrt_d_not_rat by fastforce+
+ ultimately show False
+ by simp
+qed
+
+definition S where "S = {(t::int,u::int). \<bar>t - u * sqrt d\<bar> < 1 / \<bar>u\<bar> \<and> u > 0 \<and> t > 0}"
+
+lemma S_not_empty: "S \<noteq> {}"
+proof -
+ from always_smaller[of 1] obtain t u::int
+ where "0 < t" "0 < u" "\<bar>real_of_int t - real_of_int u * sqrt d\<bar> < 1 / real 1" "1 / real 1 \<le> 1 / real_of_int \<bar>u\<bar>"
+ by auto
+ then have "\<bar>real_of_int t - real_of_int u * sqrt d\<bar> < 1 / real_of_int \<bar>u\<bar>"
+ by linarith
+ then have "(t,u) \<in> S"
+ unfolding S_def using `u > 0` `t > 0` by simp
+ then show "S \<noteq> {}"
+ by blast
+qed
+
+lemma S_infinite: "infinite S"
+proof(rule ccontr)
+ assume "\<not> infinite S"
+ then have assms: "finite S"
+ by simp
+ define tus where "tus = ((\<lambda>(t,u) . \<bar>t - u * sqrt d\<bar>) ` S)"
+ define min_tu where "min_tu = Min tus"
+ from S_not_empty have "tus \<noteq> {}"
+ unfolding tus_def by fast
+ moreover have "finite tus"
+ unfolding tus_def using `finite S` by blast
+ moreover have "\<And>x. x \<in> tus \<Longrightarrow> x > 0"
+ unfolding tus_def S_def apply(auto)
+ using not_Rats_mult
+ by (metis Rats_of_int less_int_code(1) of_int_0_eq_iff sqrt_d_not_rat)
+ ultimately have "min_tu > 0"
+ unfolding min_tu_def by simp
+ then have "\<exists>M::nat. M > 0 \<and> 1 / M < min_tu"
+ by (meson nat_approx_posE of_nat_0_less_iff zero_less_Suc)
+ then obtain M::nat where "M > 0" "1 / M < min_tu"
+ by force
+ with always_smaller obtain t u::int where "\<bar>t - u * sqrt d\<bar> < 1 / M" "1 / M \<le> 1 / \<bar>u\<bar>" "u > 0" "t > 0"
+ by metis
+ then have "(t,u) \<in> S"
+ unfolding S_def by fastforce
+ then have "\<bar>t - u * sqrt d\<bar> \<in> tus"
+ unfolding tus_def by fast
+ then have "min_tu \<le> \<bar>t - u * sqrt d\<bar>"
+ unfolding min_tu_def using `finite tus` `tus \<noteq> {}` by fastforce
+ moreover have "\<bar>t - u * sqrt d\<bar> < min_tu"
+ using `1 / M < min_tu` `\<bar>t - u * sqrt d\<bar> < 1 / M` by simp
+ ultimately show False
+ by simp
+qed
+
+lemma S_smaller: assumes "(t,u) \<in> S"
+ shows "\<bar>t\<^sup>2 - d * u\<^sup>2\<bar> < 1 + 2 * sqrt d"
+proof -
+ have 0: "\<bar>t - u * sqrt d\<bar> < 1 / \<bar>u\<bar>"
+ using assms unfolding S_def by force
+ have 3: "\<And>x. real_of_int (x\<^sup>2) = (real_of_int x)\<^sup>2"
+ by simp
+ have "\<bar>t + u * sqrt d\<bar> = \<bar>t - u * sqrt d + 2 * u * sqrt d\<bar>"
+ by fastforce
+ also have "\<dots> \<le> \<bar>t - u * sqrt d\<bar> + \<bar>2 * u * sqrt d\<bar>"
+ by simp
+ also have "\<dots> < 1 / \<bar>u\<bar> + \<bar>2 * u * sqrt d\<bar>"
+ using assms unfolding S_def by force
+ also have "\<dots> = 1 / \<bar>u\<bar> + 2 * \<bar>u\<bar> * sqrt d"
+ by (simp add: abs_mult d_greater_zero less_eq_real_def)
+ finally have 1: "\<bar>t + u * sqrt d\<bar> < 1 / \<bar>u\<bar> + 2 * \<bar>u\<bar> * sqrt d"
+ .
+ have "t\<^sup>2 - d * u\<^sup>2 = (t - sqrt d * u) * (t + sqrt d * u)"
+ apply(subst left_diff_distrib)
+ apply(subst ring_distribs)
+ apply(subst ring_distribs)
+ apply(simp)
+ by (simp add: d_greater_zero less_eq_real_def power2_eq_square)
+ then have "\<bar>t\<^sup>2 - d * u\<^sup>2\<bar> = \<bar>(t - sqrt d * u) * (t + sqrt d * u)\<bar>"
+ by simp
+ also have "\<dots> = \<bar>(t - sqrt d * u)\<bar> * \<bar>(t + sqrt d * u)\<bar>"
+ using abs_mult by blast
+ also have "\<dots> \<le> 1 / \<bar>u\<bar> * \<bar>t + sqrt d * u\<bar>"
+ apply(rule mult_right_mono)
+ using 0 apply(argo)
+ by simp
+ also have "\<dots> < 1 / \<bar>u\<bar> * (1 / \<bar>u\<bar> + 2 * \<bar>u\<bar> * sqrt d)"
+ apply(rule mult_strict_left_mono)
+ using 1 apply argo
+ using assms unfolding S_def by fastforce
+ also have "\<dots> = 1 / \<bar>u\<bar>\<^sup>2 + 2 * sqrt d"
+ apply(auto simp add: algebra_simps)
+ using assms unfolding S_def apply(simp)
+ by (simp add: power2_eq_square)
+ also have "\<dots> \<le> 1 + 2 * sqrt d"
+ apply(simp)
+ using assms unfolding S_def
+ by (clarsimp)
+ finally show ?thesis
+ by simp
+qed
+
+
+
+lemma of_int_mod_0_in_Ints: "(k::int) mod i = 0 \<Longrightarrow> k / i \<in> \<int>"
+ apply(subst (asm) dvd_eq_mod_eq_0[symmetric])
+ using of_int_divide_in_Ints by blast
+
+lemma "finite A \<Longrightarrow> \<forall>a \<in> A. finite (f a) \<Longrightarrow> finite (UNION A f)"
+ by fast
+
+
+lemma abs_set_floor: fixes y::real
+ shows "{x::int. \<bar>x\<bar> < y} \<subseteq> {-floor y..floor y}"
+ apply(auto) by linarith+
+
+lemma solutions_exists: "\<exists>x::int. \<exists>y::int. y \<noteq> 0 \<and> x\<^sup>2 - d * y\<^sup>2 = 1"
+proof -
+ have "\<exists>k::int. \<bar>k\<bar> < 1 + 2 * sqrt d \<and> infinite {(t::int, u::int). \<bar>t\<^sup>2 - d* u^2\<bar> = k \<and> u > 0 \<and> t > 0}"
+ proof(rule ccontr)
+ assume "\<nexists>x. \<bar>real_of_int x\<bar> < 1 + 2 * sqrt d \<and>
+ infinite
+ {(xb, xa). \<bar>(real_of_int xb)\<^sup>2 - d * (real_of_int xa)\<^sup>2\<bar> = real_of_int x \<and> real_of_int xa > 0 \<and> real_of_int xb > 0}"
+ then have 1: "\<And>x::int. \<bar>x\<bar> < 1 + 2 * sqrt d \<Longrightarrow>
+ finite {(t::int, u::int). \<bar>t\<^sup>2 - d * u\<^sup>2\<bar> = x \<and> u > 0 \<and> t > 0}"
+ by simp
+ have f: "finite (UNION {0..< floor (1 + 2 * sqrt d)} (\<lambda>i. {(t::int, u::int). \<bar>t\<^sup>2 - d * u\<^sup>2\<bar> = i \<and> u > 0 \<and> t > 0}))"
+ apply(rule)
+ apply(blast)
+ apply(rule 1)
+ by (auto simp add: floor_less_cancel)
+ define S' where "S' = {(xb, xa). \<bar>(real_of_int xb)\<^sup>2 - d * (real_of_int xa)\<^sup>2\<bar> < 1 + 2 * sqrt d \<and> xa > 0 \<and> xb > 0}"
+ have b: "S' \<subseteq> UNION {k::int. \<bar>k\<bar> < 1 + 2 * sqrt d} (\<lambda>i. {(t::int, u::int). \<bar>t\<^sup>2 - d * u\<^sup>2\<bar> = i \<and> u > 0 \<and> t > 0})"
+ unfolding S'_def apply(auto)
+ proof (goal_cases)
+ case (1 a b)
+ define k where "k = \<bar>(real_of_int a)\<^sup>2 - d * (real_of_int b)\<^sup>2\<bar>"
+ have "k \<in> \<int>"
+ unfolding k_def
+ using d_nat by (auto simp add: Nats_altdef2 intro!: Ints_diff Ints_mult)
+ then have "\<exists>k'::int. k' = k"
+ using Ints_cases by auto
+ then obtain k' where "k = real_of_int k'"
+ by blast
+ have "k < 1 + 2 * sqrt d"
+ using 1 k_def by blast
+ then show ?case
+ using `k = real_of_int k'` 1 k_def by simp
+ qed
+ have "S \<subseteq> S'"
+ unfolding S'_def
+ apply(rule)
+ apply(simp) using S_smaller apply(simp add: case_prod_beta') apply(rule)
+ unfolding S_def by auto
+ (* Weird behaviour with real_of_int *)
+ then have "infinite S'"
+ apply(rule infinite_super) using S_infinite .
+ moreover have "finite S'"
+ apply(rule finite_subset)
+ apply(rule b)
+ apply(rule)
+ apply(auto)
+ apply(rule rev_finite_subset[OF _ abs_set_floor])
+ apply(blast)
+ using 1 by(presburger)
+ ultimately show False by blast
+ qed
+ then obtain k::int where
+ k_def: "\<bar>k\<bar> < 1 + 2 * sqrt d" "infinite {(t::int, u::int). \<bar>t\<^sup>2 - d* u^2\<bar> = k \<and> u > 0 \<and> t > 0}"
+ by auto
+ have "{(t::int, u::int). \<bar>t\<^sup>2 - d* u^2\<bar> = k \<and> u > 0 \<and> t > 0} \<subseteq>
+ {(t::int, u::int). t\<^sup>2 - d* u^2 = k \<and> u > 0 \<and> t > 0} \<union> {(t::int, u::int). (t\<^sup>2 - d* u^2) = - k \<and> u > 0 \<and> t > 0}"
+ by force
+ then have "infinite ({(t::int, u::int). t\<^sup>2 - d* u^2 = k \<and> u > 0 \<and> t > 0} \<union> {(t::int, u::int). (t\<^sup>2 - d* u^2) = - k \<and> u > 0 \<and> t > 0})"
+ using infinite_super k_def by auto
+ then consider "infinite {(t::int, u::int). t\<^sup>2 - d* u^2 = k \<and> u > 0 \<and> t > 0}" | "infinite {(t::int, u::int). (t\<^sup>2 - d* u^2) = - k \<and> u > 0 \<and> t > 0}"
+ using infinite_Un by fast
+ then have "\<exists>k::int. \<bar>k\<bar> < 1 + 2 * sqrt d \<and> infinite {(t::int, u::int). t\<^sup>2 - d* u^2 = k \<and> u > 0 \<and> t > 0}"
+ apply(cases)
+ using k_def apply(force)
+ apply(rule)
+ apply(rule)
+ defer
+ apply(blast)
+ using k_def by auto
+ then obtain k::int where k: "\<bar>k\<bar> < 1 + 2 * sqrt d" and k_def: "infinite {(t::int, u::int). t\<^sup>2 - d * u^2 = k \<and> u > 0 \<and> t > 0}"
+ by auto
+ define Sk where "Sk = {(t::int, u::int). t\<^sup>2 - d * u^2 = k \<and> u > 0 \<and> t > 0}"
+ have "infinite Sk"
+ unfolding Sk_def using k_def by simp
+ then have "Sk \<noteq> {}"
+ unfolding Sk_def by (metis finite.emptyI)
+ have "k \<noteq> 0"
+ proof (rule ccontr)
+ assume "\<not>k \<noteq> 0"
+ then have "{(t::int, u::int). t\<^sup>2 - d * u^2 = k \<and> u > 0 \<and> t > 0} = {}"
+ using pell_not_zero by auto
+ then show False
+ using `Sk \<noteq> {}` unfolding Sk_def by simp
+ qed
+ from `k \<noteq> 0` have "(\<lambda>(t,u). (t mod k, u mod k)) ` Sk \<subseteq> ({min k (-k)..max k (-k)} \<times> {min k (-k)..max k (-k)})"
+ apply(auto simp add: min_def max_def)
+ apply (meson antisym linear neg_mod_conj not_le)
+ apply (meson linorder_neqE_linordered_idom neg_0_le_iff_le neg_mod_conj not_le order_trans)
+ apply (meson antisym linear neg_mod_bound not_less)
+ apply (meson linorder_neqE_linordered_idom neg_0_le_iff_le neg_mod_conj not_less order_trans)
+ apply (metis (full_types) linear neg_equal_0_iff_equal neg_le_iff_le not_le order_trans pos_mod_conj)
+ apply (meson linear not_less pos_mod_conj)
+ apply (metis linear neg_equal_zero neg_le_iff_le not_less order_trans pos_mod_conj)
+ by (meson le_less not_less pos_mod_conj)
+ from rev_finite_subset[OF _ this] have "finite ((\<lambda>(t,u). (t mod k, u mod k)) ` Sk)"
+ by blast
+ note pigeonhole_infinite[OF `infinite Sk` this]
+ from this have "\<exists>a0\<in>Sk. infinite {a \<in> Sk. (fst a mod k, snd a mod k) = (fst a0 mod k, snd a0 mod k)}"
+ unfolding vimage_def case_prod_beta by (simp)
+ then obtain t1 u1 where "(t1, u1) \<in> Sk" "infinite {a \<in> Sk. (fst a mod k, snd a mod k) = (t1 mod k, u1 mod k)}"
+ by force
+ define Smod where "Smod = {a \<in> Sk. (fst a mod k, snd a mod k) = (t1 mod k, u1 mod k)}"
+ have "infinite Smod"
+ unfolding Smod_def using `infinite {a \<in> Sk. (fst a mod k, snd a mod k) = (t1 mod k, u1 mod k)}` by blast
+ have "\<And>t2 u2. (t2, u2) \<in> Sk \<Longrightarrow> u2^2 = ((real_of_int t2)^2 - k) / d"
+ unfolding Sk_def using d_greater_zero
+ proof -
+ fix t2 :: int and u2 :: int
+ assume "(t2, u2) \<in> {(xa, x). (real_of_int xa)\<^sup>2 - d * (real_of_int x)\<^sup>2 = real_of_int k \<and> real_of_int x > 0 \<and> real_of_int xa > 0}"
+ then have "- 1 * real_of_int k + (real_of_int t2)\<^sup>2 = d * (real_of_int u2)\<^sup>2"
+ by force
+ then show "(real_of_int u2)\<^sup>2 = ((real_of_int t2)\<^sup>2 - real_of_int k) / d"
+ using d_greater_zero by force
+ qed
+ then have abs_u2: "\<And>t2 u2. (t2, u2) \<in> Sk \<Longrightarrow> \<bar>u2\<bar> = sqrt (((real_of_int t2)^2 - k) / d)"
+ by (metis real_sqrt_abs)
+ then have u2_eq: "\<And>t2 u2. (t2, u2) \<in> Sk \<Longrightarrow> u2 = sqrt (((real_of_int t2)^2 - k) / d) \<or> u2 = -sqrt (((real_of_int t2)^2 - k) / d) "
+ by force
+ then have u2_eq_sm: "\<And>t2 u2. (t2, u2) \<in> Smod \<Longrightarrow> u2 = sqrt (((real_of_int t2)^2 - k) / d) \<or> u2 = -sqrt (((real_of_int t2)^2 - k) / d) "
+ unfolding Smod_def by force
+ have "\<And>t2 u2. (t2, u2) \<in> Sk \<Longrightarrow> \<bar>t2\<bar> = sqrt (k + d * u2^2)"
+ unfolding Sk_def
+ proof - (* Holy shit, sledgehammer *)
+ fix t2 :: int and u2 :: int
+ assume "(t2, u2) \<in> {(xa, x). (real_of_int xa)\<^sup>2 - d * (real_of_int x)\<^sup>2 = real_of_int k \<and> real_of_int x > 0 \<and> real_of_int xa > 0}"
+ then have "real_of_int t2 * real_of_int t2 - d * (real_of_int u2 * real_of_int u2) = real_of_int k"
+ by (simp add: power2_eq_square)
+ then show "\<bar>real_of_int t2\<bar> = sqrt (real_of_int k + d * (real_of_int u2)\<^sup>2)"
+ by (metis diff_add_cancel power2_eq_square real_sqrt_abs2)
+ qed
+ then have "\<And>t2 u2. (t2, u2) \<in> Sk \<Longrightarrow> t2 = sqrt (k + d * u2^2) \<or> t2 = -sqrt (k + d * u2^2)"
+ by force
+ then have 431234: "\<And>t2 u2. (t2, u2) \<in> Smod \<Longrightarrow> t2 = sqrt (k + d * u2^2) \<or> t2 = -sqrt (k + d * u2^2)"
+ unfolding Smod_def by force
+ have "\<exists>t2 u2. (t2, u2) \<in> Smod \<and> \<bar>t1\<bar> \<noteq> \<bar>t2\<bar>"
+ proof(rule ccontr)
+ assume "\<nexists>t2 u2. (t2, u2) \<in> Smod \<and> \<bar>t1\<bar> \<noteq> \<bar>t2\<bar>"
+ then have a: "\<And>t2 u2. (t2, u2) \<in> Smod \<Longrightarrow> \<bar>t1\<bar> = \<bar>t2\<bar>"
+ by blast
+ then have a: "fst ` Smod \<subseteq> {x. \<bar>x\<bar> = \<bar>t1\<bar>}"
+ by (auto)
+ have "finite {x. \<bar>x\<bar> = \<bar>t1\<bar>}"
+ by (metis (mono_tags, lifting) infinite_int_iff_unbounded mem_Collect_eq order_less_irrefl)
+ note finite_subset[OF a this]
+ from pigeonhole_infinite[OF `infinite Smod` this]
+ obtain a0 where l: "a0\<in>Smod" "infinite {a \<in> Smod. fst a = fst a0}"
+ by blast
+ have 3: "(\<lambda>(a,b). (real_of_int a, real_of_int b)) ` {a \<in> Smod. fst a = fst a0} \<subseteq> {fst a0} \<times> {sqrt (((real_of_int (fst a0))^2 - k) / d), -sqrt (((real_of_int (fst a0))^2 - k) / d)}"
+ apply(auto)
+ apply(drule u2_eq_sm)
+ by blast
+ have 4: "infinite ((\<lambda>(a,b). (real_of_int a, real_of_int b)) ` {a \<in> Smod. fst a = fst a0})"
+ apply(rule)
+ apply(subst (asm) finite_image_iff)
+ apply(rule)
+ using `infinite {a \<in> Smod. fst a = fst a0}` by (auto)
+ have "{a \<in> Smod. fst a = fst a0} = {(a,b) \<in> Smod. a = fst a0}"
+ by fastforce
+ have "finite ( {real_of_int (fst a0)} \<times> {sqrt (((real_of_int (fst a0))^2 - k) / d), -sqrt (((real_of_int (fst a0))^2 - k) / d)})"
+ by blast
+ moreover have "infinite ({real_of_int (fst a0)} \<times> {sqrt (((real_of_int (fst a0))^2 - k) / d), -sqrt (((real_of_int (fst a0))^2 - k) / d)})"
+ apply(rule infinite_super)
+ apply(rule 3)
+ by (rule 4)
+ ultimately show False by blast
+ qed
+ then obtain t2 u2 where "(t2, u2) \<in> Smod" "\<bar>t1\<bar> \<noteq> \<bar>t2\<bar>"
+ by blast
+ have "\<bar>real_of_int u1\<bar> \<noteq> \<bar>real_of_int u2\<bar>"
+ apply(subst abs_u2)
+ using `(t1, u1) \<in> Sk` apply(simp)
+ apply(subst abs_u2)
+ using `(t2, u2) \<in> Smod` unfolding Smod_def apply(blast)
+ apply(subst real_sqrt_eq_iff)
+ apply(subst divide_cancel_right)
+ using d_greater_zero apply(simp)
+ apply(rule ccontr)
+ apply(simp)
+ using `\<bar>t1\<bar> \<noteq> \<bar>t2\<bar>`
+ by (metis of_int_abs of_int_eq_iff real_sqrt_abs)
+ then have "\<bar>u1\<bar> \<noteq> \<bar>u2\<bar>"
+ by auto
+ obtain D::int where D: "D = d"
+ using Nats_cases d_nat of_int_of_nat_eq by blast
+ have "k * k = (t1\<^sup>2 - d* u1^2) * (t2\<^sup>2 - d* u2^2)"
+ using `(t1,u1) \<in> Sk` `(t2, u2) \<in> Smod`
+ unfolding Smod_def Sk_def by (auto)
+ also have "\<dots> = (t1*t2 - d*u1*u2)\<^sup>2 - d*(t1*u2 - t2 * u1)^2"
+ by (auto simp add: algebra_simps power2_diff power_mult_distrib power2_eq_square)
+ finally have "k * k = (t1*t2 - d*u1*u2)\<^sup>2 - d*(t1*u2 - t2 * u1)^2" .
+ then have "k * k / (k * k) = ((t1*t2 - d*u1*u2)\<^sup>2 - d*(t1*u2 - t2 * u1)^2) / (k * k)"
+ apply(subst divide_cancel_right)
+ using `k \<noteq> 0` by linarith
+ then have "1 = ((t1*t2 - d*u1*u2)\<^sup>2 - d*(t1*u2 - t2 * u1)^2) / (k * k)"
+ using `k \<noteq> 0` by force
+ also have "\<dots> = (t1*t2 - d*u1*u2)\<^sup>2 / (k * k) - d*(t1*u2 - t2 * u1)^2 / (k * k)"
+ by algebra
+ also have "\<dots> = (t1*t2 - d*u1*u2)\<^sup>2 / k\<^sup>2 - d*(t1*u2 - t2 * u1)^2 / k^2"
+ by (simp add: power2_eq_square)
+ also have "\<dots> = ((t1*t2 - d*u1*u2) / k)\<^sup>2 - d*((t1*u2 - t2 * u1) / k)^2"
+ apply(subst power_divide)
+ apply(subst power_divide)
+ by simp
+ finally have 1: "1 = ((t1*t2 - d*u1*u2) / k)\<^sup>2 - d*((t1*u2 - t2 * u1) / k)^2"
+ by blast
+ have "(real_of_int (t1 * t2) - d * real_of_int u1 * real_of_int u2) = real_of_int (t1 * t2 - D * u1 * u2)"
+ apply(subst D[symmetric])
+ by simp
+ have "real_of_int (t1 * u2 - t2 * u1) = real_of_int (t1 * u2 - t2 * u1)"
+ by simp
+ have b: "t2 mod k = t1 mod k"
+ using `(t2, u2) \<in> Smod` unfolding Smod_def by force
+ have c: "u2 mod k = u1 mod k"
+ using `(t2, u2) \<in> Smod` unfolding Smod_def by force
+ have lll: "(t1 * t1 - D * u1 * u1) = k"
+ using `(t1, u1) \<in> Sk` D[symmetric] unfolding Sk_def
+ apply(auto simp add: power2_eq_square)
+ by (metis floor_of_int mult.assoc of_int_diff of_int_mult)
+ have "(t1 * t2 - D * u1 * u2) mod k = 0"
+ apply(subst mod_diff_left_eq[symmetric])
+ apply(subst mod_diff_right_eq[symmetric])
+ apply(subst mod_mult_right_eq[symmetric])
+ apply(subst b)
+ apply(subst mod_mult_right_eq)
+ apply(subst (2) mod_mult_right_eq[symmetric])
+ apply(subst c)
+ by (simp add: mod_simps lll)
+ then have new_t: "real_of_int (t1 * t2 - D * u1 * u2) / real_of_int k \<in> \<int>"
+ using of_int_mod_0_in_Ints by metis
+ have us_greater_zero : "u2 > 0" "u1 > 0"
+ using `(t1, u1) \<in> Sk` `(t2, u2) \<in> Smod` unfolding Smod_def Sk_def by auto
+ have ts_greater_zero: "t1 > 0" "t2 > 0"
+ using `(t1,u1) \<in> Sk` `(t2, u2) \<in> Smod` unfolding Smod_def Sk_def by auto
+ have a: "(t1 * u2 - t2 * u1) mod k = 0"
+ apply(subst mod_diff_left_eq[symmetric])
+ apply(subst mod_mult_right_eq[symmetric])
+ apply(subst c)
+ apply(subst mod_diff_right_eq[symmetric])
+ apply(subst (2) mod_mult_left_eq[symmetric])
+ apply(subst b)
+ by (simp add: mod_simps)
+ then have new_u: "real_of_int (t1 * u2 - t2 * u1) / real_of_int k \<in> \<int>"
+ using of_int_mod_0_in_Ints by metis
+ have olong_johnson: "t1 * u2 - t2 * u1 \<noteq> 0"
+ proof (rule ccontr)
+ assume "\<not> t1 * u2 - t2 * u1 \<noteq> 0"
+ then have a: "t1 * u2 = t2 * u1"
+ by simp
+ from a have t1: "t1 = (t2 * u1) / u2"
+ using us_greater_zero
+ by (metis eq_divide_eq floor_of_int less_irrefl of_int_0 of_int_mult)
+ from a have u1: "u1 = (t1 * u2) / t2"
+ using ts_greater_zero by auto
+ from a have fraci: "t1 / t2 = u1 / u2"
+ using ts_greater_zero us_greater_zero
+ by (metis divide_divide_eq_left mult.commute t1 u1)
+ from a have "1 = ((t1*t2 - d*u1*u2) / k)\<^sup>2"
+ using `1 = ((t1*t2 - d*u1*u2) / k)\<^sup>2 - d*((t1*u2 - t2 * u1) / k)^2` by simp
+ also have "\<dots> = ((t1*t2 - d*u1*u2)^2 / k^2)"
+ by (simp add: power2_eq_square)
+ finally have "k^2 = (t1 * t2 - d * u1 * u2)\<^sup>2"
+ by simp
+ then have "\<bar>k\<bar> = \<bar>t1 * t2 - d * u1 * u2\<bar>"
+ by (metis of_int_abs of_int_mult power2_eq_square real_sqrt_abs2)
+ then consider (plus) "k = t1 * t2 - d * u1 * u2" | (minus) "k = - (t1 * t2 - d * u1 * u2)"
+ by linarith
+ then show False
+ proof(cases)
+ case plus
+ then have "k = (t2 * u1) / u2 * t2 - d * (t1 * u2) / t2 * u2"
+ apply(subst (asm) u1) apply(subst (asm) of_int_mult) apply(subst (asm) t1)
+ by argo
+ also have "\<dots> = t2 * u1 / u2 * t2 - t1 / t2 * d * u2 * u2"
+ by force
+ also have "\<dots> = u1 / u2 * t2 * t2 - u1 / u2 * d * u2 * u2"
+ apply(subst fraci)
+ by fastforce
+ also have "\<dots> = (u1 / u2) * (real_of_int t2 * t2 - d * u2 * u2)"
+ by (simp add: algebra_simps)
+ finally have "k = (u1 / u2) * (real_of_int t2 * t2 - d * u2 * u2)"
+ by simp
+ moreover have "(real_of_int t2 * t2 - d * u2 * u2) = k"
+ using `(t2, u2) \<in> Smod` unfolding Smod_def Sk_def by (simp add: power2_eq_square)
+ ultimately have "k = (u1 / u2) * k"
+ by simp
+ then have "1 = (u1 / u2)"
+ using `k \<noteq> 0`
+ by (metis mult_cancel_right1 of_int_eq_0_iff)
+ then have "u1 = u2"
+ by simp
+ then show False using `\<bar>u1\<bar> \<noteq> \<bar>u2\<bar>` by blast
+ next
+ case minus
+ then have "k = - ((t2 * u1) / u2 * t2 - d * (t1 * u2) / t2 * u2)"
+ apply(subst (asm) u1) apply(subst (asm) of_int_mult) apply(subst (asm) t1)
+ by argo
+ also have "\<dots> = - (t2 * u1 / u2 * t2 - t1 / t2 * d * u2 * u2)"
+ by force
+ also have "\<dots> = -( u1 / u2 * t2 * t2 - u1 / u2 * d * u2 * u2)"
+ apply(subst fraci)
+ by fastforce
+ also have "\<dots> = -( (u1 / u2) * (real_of_int t2 * t2 - d * u2 * u2))"
+ by (simp add: algebra_simps)
+ finally have "- k = (u1 / u2) * (real_of_int t2 * t2 - d * u2 * u2)"
+ by simp
+ moreover have "(real_of_int t2 * t2 - d * u2 * u2) = k"
+ using `(t2, u2) \<in> Smod` unfolding Smod_def Sk_def by (simp add: power2_eq_square)
+ ultimately have "- k = (u1 / u2) * k"
+ by simp
+ then have "- 1 = (u1 / u2)"
+ using `k \<noteq> 0`
+ by (metis add.inverse_inverse floor_of_int mult.commute mult_cancel_left2 mult_minus_right of_int_0 of_int_minus)
+ then show False
+ by (metis \<open>\<bar>real_of_int u1\<bar> \<noteq> \<bar>real_of_int u2\<bar>\<close> abs_minus_cancel divide_eq_minus_1_iff)
+ qed
+ qed
+ define new_u where "new_u = real_of_int (t1 * u2 - t2 * u1) / real_of_int k"
+ define new_t where "new_t = real_of_int (t1 * t2 - D * u1 * u2) / real_of_int k"
+ thm 1 new_t new_u
+ have "new_u \<in> \<int>"
+ unfolding new_u_def using new_u by auto
+ have "new_t \<in> \<int>"
+ unfolding new_t_def using new_t by auto
+ have "new_u \<noteq> 0"
+ unfolding new_u_def using olong_johnson `k \<noteq> 0`
+ using divide_eq_0_iff of_int_eq_0_iff by blast
+ have 1: "new_t^2 - d * new_u^2 = 1"
+ unfolding new_t_def new_u_def using D[symmetric] 1 by simp
+ obtain t::int where t_def: "t = new_t"
+ using Ints_cases `new_t \<in> \<int>` by metis
+ obtain u::int where u_def: "u = new_u"
+ using Ints_cases `new_u \<in> \<int>` by metis
+ have g: "(real_of_int t)^2 - d * (real_of_int u)^2 = 1"
+ using u_def t_def 1 by simp
+ have h: "real_of_int u \<noteq> 0"
+ using `new_u \<noteq> 0` u_def by blast
+ then show ?thesis
+ using g h by blast
+qed
+
+ end
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Haslbeck_Maximilian_max.haslbeck@mytum.de_640/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,2 @@
+Using temporary directory '/tmp/tmp.YsLPdkMbQ7'
+Exercise path '/home/admin/suites/exbeoriginal' not found
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Hellauer_Fabian_f.hellauer@tum.de_656/FDS_hw.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,505 @@
+(*\<Rightarrow> Tobias*)
+theory FDS_hw
+ imports
+ Complex_Main
+ "~~/src/HOL/Data_Structures/Tree2"
+ "~~/src/HOL/Data_Structures/Cmp"
+ "~~/src/HOL/Data_Structures/Isin2"
+begin
+
+(* From email:
+
+Hallo Peter,
+
+vieles davon ist aus Thys/RBT.thy und Thys/RBT_Set.thy kopiert, zum Teil
+aber modifiziert und ziemlich verstreut, deswegen hab ich das Markieren
+der Kopien irgendwann aufgegeben.
+
+Neu ist auf jeden Fall:
+
+- die Formulierung mit locales, damit man für verschieden Erweiterungen
+die Beweise nicht wiederholen muss
+
+- das extension field mitsamt der Invariante, die bei insert/delete
+erhalten bleibt
+
+- die funktionen rank und select und ihre Korrektheitsbeweise
+
+Viele Grüße,
+
+Fabian
+
+*)
+
+datatype color = Red | Black
+
+type_synonym ('a,'x) xrbt = "('a,color*'x)tree" (*for convenience, store the extension value together with the color*)
+locale xRBT = (*extended RBT*)
+ fixes x_leaf :: "'x" (*extension value which should be assigned to Leafs*)
+ fixes f :: "'x\<Rightarrow>'a\<Rightarrow>'x\<Rightarrow>'x" (*Todo: pass the whole subtrees instead of the x-values?*)
+ and t_f :: "'x\<Rightarrow>'a\<Rightarrow>'x\<Rightarrow>nat"
+assumes t_f_constant: "\<exists>c.\<forall>x. t_f x \<le> c" (*if t_f is constant it doesn't change the complexity class of ins, del,...*)
+begin
+
+ fun get_x :: "('a,'x) xrbt \<Rightarrow> 'x" where
+ "get_x Leaf = x_leaf" |
+ "get_x (Node (_,x) _ _ _) = x"
+
+ abbreviation R where "R x l a r \<equiv> Node (Red,x) l a r"
+ abbreviation B where "B x l a r \<equiv> Node (Black,x) l a r"
+
+ (*smart constructors*)
+ abbreviation make_R where "make_R l a r \<equiv> Node (Red,f (get_x l) a (get_x r)) l a r"
+ abbreviation make_B where "make_B l a r \<equiv> Node (Black,f (get_x l) a (get_x r)) l a r"
+
+ fun baliL :: "('a,'x) xrbt \<Rightarrow> 'a \<Rightarrow> ('a,'x) xrbt \<Rightarrow> ('a,'x) xrbt" where
+ "baliL (R x2 (R x1 t1 a1 t2) a2 t3) a3 t4 = make_R (B x1 t1 a1 t2) a2 (make_B t3 a3 t4)" --\<open>No need to recompute x1\<close> |
+ "baliL (R x1 t1 a1 (R x2 t2 a2 t3)) a3 t4 = make_R (make_B t1 a1 t2) a2 (make_B t3 a3 t4)" |
+ "baliL t1 a t2 = make_B t1 a t2"
+
+ fun baliR :: "('a,'x) xrbt \<Rightarrow> 'a \<Rightarrow> ('a,'x) xrbt \<Rightarrow> ('a,'x) xrbt" where
+ "baliR t1 a1 (R x3 (R x2 t2 a2 t3) a3 t4) = make_R (make_B t1 a1 t2) a2 (make_B t3 a3 t4)" |
+ "baliR t1 a1 (R x2 t2 a2 (R x3 t3 a3 t4)) = make_R (make_B t1 a1 t2) a2 (B x3 t3 a3 t4)" |
+ "baliR t1 a t2 = make_B t1 a t2"
+
+ fun paint :: "color \<Rightarrow> ('a,'x) xrbt \<Rightarrow> ('a,'x) xrbt" where
+ "paint c Leaf = Leaf" |
+ "paint c (Node (_,x) l a r) = Node (c,x) l a r"
+
+ fun baldL :: "('a,'x) xrbt \<Rightarrow> 'a \<Rightarrow> ('a,'x) xrbt \<Rightarrow> ('a,'x) xrbt" where
+ "baldL (R x1 t1 a1 t2) a2 t3 = make_R (B x1 t1 a1 t2) a2 t3" |
+ "baldL bl a1 (B x2 t1 a2 t2) = baliR bl a1 (R x2 t1 a2 t2)" |
+ "baldL bl a1 (R x3 (B x2 t1 a2 t2) a3 t3) = make_R (make_B bl a1 t1) a2 (baliR t2 a3 (paint Red t3))" |
+ "baldL t1 a t2 = make_R t1 a t2"
+
+ fun baldR :: "('a,'x) xrbt \<Rightarrow> 'a \<Rightarrow> ('a,'x) xrbt \<Rightarrow> ('a,'x) xrbt" where
+ "baldR t1 x (R b t2 y t3) = make_R t1 x (B b t2 y t3)" |
+ "baldR (B a t1 x t2) y t3 = baliL (R a t1 x t2) y t3" |
+ "baldR (R a t1 x (B b t2 y t3)) z t4 = make_R (baliL (paint Red t1) x t2) y (make_B t3 z t4)" |
+ "baldR t1 x t2 = make_R t1 x t2"
+
+ fun combine :: "('a,'x) xrbt \<Rightarrow> ('a,'x) xrbt \<Rightarrow> ('a,'x) xrbt" where
+ "combine Leaf t = t" |
+ "combine t Leaf = t" |
+ "combine (R x1 t1 a t2) (R x3 t3 c t4) =
+ (case combine t2 t3 of
+ R x2 u2 b u3 \<Rightarrow> (make_R (make_R t1 a u2) b (make_R u3 c t4)) |
+ t23 \<Rightarrow> make_R t1 a (make_R t23 c t4))" |
+ "combine (B x1 t1 a t2) (B x3 t3 c t4) =
+ (case combine t2 t3 of
+ R x2 t2' b t3' \<Rightarrow> make_R (make_B t1 a t2') b (make_B t3' c t4) |
+ t23 \<Rightarrow> baldL t1 a (make_B t23 c t4))" |
+ "combine t1 (R x1 t2 a t3) = make_R (combine t1 t2) a t3" |
+ "combine (R x1 t1 a t2) t3 = make_R t1 a (combine t2 t3)"
+
+end
+
+section \<open>Red-Black Tree Implementation of Sets\<close>
+locale xRBT_ins_del = xRBT _ f for f :: "'x\<Rightarrow>'a::linorder\<Rightarrow>'x\<Rightarrow>'x"
+begin
+ fun ins :: "'a \<Rightarrow> ('a,'x) xrbt \<Rightarrow> ('a,'x) xrbt" where
+ "ins x Leaf = make_R Leaf x Leaf" |
+ "ins x (B x1 l a r) =
+ (case cmp x a of
+ LT \<Rightarrow> baliL (ins x l) a r |
+ GT \<Rightarrow> baliR l a (ins x r) |
+ EQ \<Rightarrow> B x1 l a r)" |
+ "ins x (R x1 l a r) =
+ (case cmp x a of
+ LT \<Rightarrow> make_R (ins x l) a r |
+ GT \<Rightarrow> make_R l a (ins x r) |
+ EQ \<Rightarrow> R x1 l a r)"
+
+ definition insert :: "'a \<Rightarrow> ('a,'x) xrbt \<Rightarrow> ('a,'x) xrbt" where
+ "insert x t = paint Black (ins x t)"
+
+ fun (in xRBT) get_color :: "('a,'x) xrbt \<Rightarrow> color" where
+ "get_color Leaf = Black" |
+ "get_color (Node (c,_) _ _ _) = c"
+
+ fun del :: "'a \<Rightarrow> ('a,'x) xrbt \<Rightarrow> ('a,'x) xrbt" where
+ "del x Leaf = Leaf" |
+ "del x (Node _ l a r) =
+ (case cmp x a of
+ LT \<Rightarrow> if l \<noteq> Leaf \<and> get_color l = Black
+ then baldL (del x l) a r else make_R (del x l) a r |
+ GT \<Rightarrow> if r \<noteq> Leaf\<and> get_color r = Black
+ then baldR l a (del x r) else make_R l a (del x r) |
+ EQ \<Rightarrow> combine l r)"
+
+ definition delete :: "'a \<Rightarrow> ('a,'x) xrbt \<Rightarrow> ('a,'x) xrbt" where
+ "delete x t = paint Black (del x t)"
+
+subsection "Functional Correctness Proofs"
+ lemma (in xRBT) inorder_paint: "inorder(paint c t) = inorder t"
+ by(cases t) (auto)
+
+ lemma (in xRBT) inorder_baliL:
+ "inorder(baliL l a r) = inorder l @ a # inorder r"
+ by(cases "(l,a,r)" rule: baliL.cases) (auto)
+
+ lemma (in xRBT) inorder_baliR:
+ "inorder(baliR l a r) = inorder l @ a # inorder r"
+ by(cases "(l,a,r)" rule: baliR.cases) (auto)
+
+ lemma inorder_ins:
+ "sorted(inorder t) \<Longrightarrow> inorder(ins x t) = ins_list x (inorder t)"
+ by(induction x t rule: ins.induct)
+ (auto simp: ins_list_simps inorder_baliL inorder_baliR)
+
+ lemma inorder_insert:
+ "sorted(inorder t) \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)"
+ by (simp add: insert_def inorder_ins inorder_paint)
+
+ lemma inorder_baldL:
+ "inorder(baldL l a r) = inorder l @ a # inorder r"
+ by(cases "(l,a,r)" rule: baldL.cases)
+ (auto simp: inorder_baliL inorder_baliR inorder_paint)
+
+ lemma inorder_baldR:
+ "inorder(baldR l a r) = inorder l @ a # inorder r"
+ by(cases "(l,a,r)" rule: baldR.cases)
+ (auto simp: inorder_baliL inorder_baliR inorder_paint)
+
+ lemma inorder_combine:
+ "inorder(combine l r) = inorder l @ inorder r"
+ by(induction l r rule: combine.induct)
+ (auto simp: inorder_baldL inorder_baldR split: tree.split color.split)
+
+ lemma inorder_del:
+ "sorted(inorder t) \<Longrightarrow> inorder(del x t) = del_list x (inorder t)"
+ by(induction x t rule: del.induct)
+ (auto simp: del_list_simps inorder_combine inorder_baldL inorder_baldR)
+
+ lemma inorder_delete:
+ "sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
+ by (auto simp: delete_def inorder_del inorder_paint)
+
+end
+
+
+text\<open>The proofs are due to Markus Reiter and Alexander Krauss.\<close>
+context xRBT begin
+fun bheight :: "('a,'x) xrbt \<Rightarrow> nat" where
+"bheight Leaf = 0" |
+"bheight (Node (c,_) l x r) = (if c = Black then bheight l + 1 else bheight l)"
+
+fun invc :: "('a,'x) xrbt \<Rightarrow> bool" where
+"invc Leaf = True" |
+"invc (Node (c,_) l a r) =
+ (invc l \<and> invc r \<and> (c = Red \<longrightarrow> get_color l = Black \<and> get_color r = Black))"
+
+fun invc2 :: "('a,'x) xrbt \<Rightarrow> bool" \<comment> \<open>Weaker version\<close> where
+"invc2 Leaf = True" |
+"invc2 (Node c l a r) = (invc l \<and> invc r)"
+
+fun invh :: "('a,'x) xrbt \<Rightarrow> bool" where
+"invh Leaf = True" |
+"invh (Node c l x r) = (invh l \<and> invh r \<and> bheight l = bheight r)"
+
+lemma invc2I: "invc t \<Longrightarrow> invc2 t"
+ apply (cases t)
+ apply simp+
+ apply (case_tac x21)
+ by simp
+
+fun invx where
+ "invx Leaf = True" |
+ "invx (Node (col,x) l a r) \<longleftrightarrow> invx l \<and> invx r \<and> x = f (get_x l) a (get_x r)"
+
+definition xrbt :: "('a,'x) xrbt \<Rightarrow> bool" where
+"xrbt t = (invx t \<and> invc t \<and> invh t \<and> get_color t = Black)"
+
+lemma color_paint_Black: "get_color (paint Black t) = Black"
+by (cases t) auto
+
+theorem xrbt_Leaf: "xrbt Leaf"
+by (simp add: xrbt_def)
+
+lemma paint_invc2: "invc2 t \<Longrightarrow> invc2 (paint c t)"
+ by (cases t) auto
+
+lemma paint_invx: "invx t \<Longrightarrow> invx (paint c t)"
+by (cases t) auto
+
+lemma invc_paint_Black: "invc2 t \<Longrightarrow> invc (paint Black t)"
+by (cases t) auto
+
+lemma invh_paint: "invh t \<Longrightarrow> invh (paint c t)"
+by (cases t) auto
+
+lemma invx_paint: "invx t \<Longrightarrow> invx (paint c t)"
+ by (cases t) auto
+
+lemma invc_baliL:
+ "\<lbrakk>invc2 l; invc r\<rbrakk> \<Longrightarrow> invc (baliL l a r)"
+by (induct l a r rule: baliL.induct) auto
+
+lemma invc_baliR:
+ "\<lbrakk>invc l; invc2 r\<rbrakk> \<Longrightarrow> invc (baliR l a r)"
+by (induct l a r rule: baliR.induct) auto
+
+lemma bheight_baliL:
+ "bheight l = bheight r \<Longrightarrow> bheight (baliL l a r) = Suc (bheight l)"
+by (induct l a r rule: baliL.induct) auto
+
+lemma bheight_baliR:
+ "bheight l = bheight r \<Longrightarrow> bheight (baliR l a r) = Suc (bheight l)"
+by (induct l a r rule: baliR.induct) auto
+
+lemma invh_baliL:
+ "\<lbrakk> invh l; invh r; bheight l = bheight r \<rbrakk> \<Longrightarrow> invh (baliL l a r)"
+by (induct l a r rule: baliL.induct) auto
+
+lemma invh_baliR:
+ "\<lbrakk> invh l; invh r; bheight l = bheight r \<rbrakk> \<Longrightarrow> invh (baliR l a r)"
+by (induct l a r rule: baliR.induct) auto
+
+lemma invx_baliL:
+ "\<lbrakk> invx l; invx r \<rbrakk> \<Longrightarrow> invx (baliL l a r)"
+by (induct l a r rule: baliL.induct) auto
+
+lemma invx_baliR:
+ "\<lbrakk> invx l; invx r \<rbrakk> \<Longrightarrow> invx (baliR l a r)"
+by (induct l a r rule: baliR.induct) auto
+
+end
+
+subsubsection \<open>Insertion\<close>
+context xRBT_ins_del
+begin
+lemma invc_ins: assumes "invc t"
+ shows "get_color t = Black \<Longrightarrow> invc (ins x t)" "invc2 (ins x t)"
+using assms
+by (induct x t rule: ins.induct) (auto simp: invc_baliL invc_baliR invc2I)
+
+lemma invh_ins: assumes "invh t"
+ shows "invh (ins x t)" "bheight (ins x t) = bheight t"
+using assms
+by(induct x t rule: ins.induct)
+ (auto simp: invh_baliL invh_baliR bheight_baliL bheight_baliR)
+
+lemma invx_ins: assumes "invx t"
+ shows "invx (ins x t)"
+using assms
+by(induct x t rule: ins.induct)
+ (auto simp: invx_baliL invx_baliR)
+
+theorem rbt_insert: "xrbt t \<Longrightarrow> xrbt (insert x t)"
+by (simp add: invc_ins(2) invh_ins(1) invx_ins color_paint_Black invc_paint_Black invh_paint invx_paint
+ xrbt_def insert_def)
+
+subsubsection \<open>Deletion\<close>
+
+lemma bheight_paint_Red:
+ "get_color t = Black \<Longrightarrow> bheight (paint Red t) = bheight t - 1"
+by (cases t) auto
+
+lemma invh_baldL_invc:
+ "\<lbrakk> invh l; invh r; bheight l + 1 = bheight r; invc r \<rbrakk>
+ \<Longrightarrow> invh (baldL l a r) \<and> bheight (baldL l a r) = bheight l + 1"
+by (induct l a r rule: baldL.induct)
+ (auto simp: invh_baliR invh_paint bheight_baliR bheight_paint_Red)
+
+lemma invh_baldL_Black:
+ "\<lbrakk> invh l; invh r; bheight l + 1 = bheight r; get_color r = Black \<rbrakk>
+ \<Longrightarrow> invh (baldL l a r) \<and> bheight (baldL l a r) = bheight r"
+by (induct l a r rule: baldL.induct) (auto simp add: invh_baliR bheight_baliR)
+
+lemma invc_baldL: "\<lbrakk>invc2 l; invc r; get_color r = Black\<rbrakk> \<Longrightarrow> invc (baldL l a r)"
+ by (induct l a r rule: baldL.induct) (simp_all add: invc_baliR)
+
+lemma invx_baldL: "\<lbrakk>invx l; invx r\<rbrakk> \<Longrightarrow> invx (baldL l a r)"
+ by (induct l a r rule: baldL.induct) (simp_all add: invx_baliR paint_invx)
+
+lemma invc2_baldL: "\<lbrakk> invc2 l; invc r \<rbrakk> \<Longrightarrow> invc2 (baldL l a r)"
+by (induct l a r rule: baldL.induct) (auto simp: invc_baliR paint_invc2 invc2I)
+
+lemma invh_baldR_invc:
+ "\<lbrakk> invh l; invh r; bheight l = bheight r + 1; invc l \<rbrakk>
+ \<Longrightarrow> invh (baldR l a r) \<and> bheight (baldR l a r) = bheight l"
+by(induct l a r rule: baldR.induct)
+ (auto simp: invh_baliL bheight_baliL invh_paint bheight_paint_Red)
+
+lemma invc_baldR: "\<lbrakk>invc a; invc2 b; get_color a = Black\<rbrakk> \<Longrightarrow> invc (baldR a x b)"
+by (induct a x b rule: baldR.induct) (simp_all add: invc_baliL)
+
+lemma invx_baldR: "\<lbrakk>invx a; invx b\<rbrakk> \<Longrightarrow> invx (baldR a x b)"
+by (induct a x b rule: baldR.induct) (simp_all add: invx_baliL paint_invx)
+
+lemma invc2_baldR: "\<lbrakk> invc l; invc2 r \<rbrakk> \<Longrightarrow>invc2 (baldR l x r)"
+by (induct l x r rule: baldR.induct) (auto simp: invc_baliL paint_invc2 invc2I)
+
+lemma invh_combine:
+ "\<lbrakk> invh l; invh r; bheight l = bheight r \<rbrakk>
+ \<Longrightarrow> invh (combine l r) \<and> bheight (combine l r) = bheight l"
+by (induct l r rule: combine.induct)
+ (auto simp: invh_baldL_Black split: tree.splits color.splits)
+
+lemma invx_combine:
+ "\<lbrakk> invx l; invx r\<rbrakk>
+ \<Longrightarrow> invx (combine l r)"
+ apply (induct l r rule: combine.induct)
+ apply
+ (auto simp: invx_baldL split: tree.splits color.splits)
+ done
+
+lemma invc_combine:
+ assumes "invc l" "invc r"
+ shows "get_color l = Black \<Longrightarrow> get_color r = Black \<Longrightarrow> invc (combine l r)"
+ "invc2 (combine l r)"
+using assms
+by (induct l r rule: combine.induct)
+ (auto simp: invc_baldL invc2I split: tree.splits color.splits)
+
+lemma neq_LeafD: "t \<noteq> Leaf \<Longrightarrow> \<exists>c l x r. t = Node c l x r"
+by(cases t) auto
+
+lemma del_invc_invh: "invh t \<Longrightarrow> invc t \<Longrightarrow> invx t \<Longrightarrow> invx (del x t) \<and> invh (del x t) \<and>
+ (get_color t = Red \<and> bheight (del x t) = bheight t \<and> invc (del x t) \<or>
+ get_color t = Black \<and> bheight (del x t) = bheight t - 1 \<and> invc2 (del x t))"
+proof (induct x t rule: del.induct)
+case (2 x c _ y)
+ have "x = y \<or> x < y \<or> x > y" by auto
+ thus ?case proof (elim disjE)
+ assume "x = y"
+ with 2 show ?thesis
+ by (cases c; case_tac a) (simp_all add: invh_combine invc_combine invx_combine)
+ next
+ assume "x < y"
+ with 2 show ?thesis
+ by(cases c; case_tac a)
+ (auto simp: invh_baldL_invc invc_baldL invc2_baldL invx_baldL dest: neq_LeafD)
+ next
+ assume "y < x"
+ with 2 show ?thesis
+ by(cases c; case_tac a)
+ (auto simp: invh_baldR_invc invc_baldR invc2_baldR invx_baldR dest: neq_LeafD)
+ qed
+qed auto
+
+theorem rbt_delete: "xrbt t \<Longrightarrow> xrbt (delete k t)"
+ by (metis color.distinct(1) color_paint_Black del_invc_invh delete_def invc_paint_Black invh_paint paint_invx xrbt_def)
+
+text \<open>Overall correctness:\<close>
+
+interpretation (*Todo: Difference to sublocale?*) Set_by_Ordered
+where empty = Leaf and isin = isin and insert = insert and delete = delete
+and inorder = inorder and inv = xrbt
+proof (standard, goal_cases)
+ case 1 show ?case by simp
+next
+ case 2 thus ?case by(simp add: isin_set)
+next
+ case 3 thus ?case by(simp add: inorder_insert)
+next
+ case 4 thus ?case by(simp add: inorder_delete)
+next
+ case 5 thus ?case by (simp add: xrbt_Leaf)
+next
+ case 6 thus ?case by (simp add: rbt_insert)
+next
+ case 7 thus ?case by (simp add: rbt_delete)
+qed
+
+
+subsection \<open>Height-Size Relation\<close>
+
+lemma neq_Black[simp]: "(c \<noteq> Black) = (c = Red)"
+by (cases c) auto
+
+lemma rbt_height_bheight_if_nat: "invc t \<Longrightarrow> invh t \<Longrightarrow>
+ height t \<le> (if get_color t = Black then 2 * bheight t else 2 * bheight t + 1)"
+by(induction t) (auto split: if_split_asm)
+
+lemma rbt_height_bheight_if: "invc t \<Longrightarrow> invh t \<Longrightarrow>
+ (if get_color t = Black then height t / 2 else (height t - 1) / 2) \<le> bheight t"
+by(induction t) (auto split: if_split_asm)
+
+lemma rbt_height_bheight: "xrbt t \<Longrightarrow> height t / 2 \<le> bheight t "
+by(auto simp: xrbt_def dest: rbt_height_bheight_if)
+
+lemma bheight_size_bound: "invc t \<Longrightarrow> invh t \<Longrightarrow> size1 t \<ge> 2 ^ (bheight t)"
+by (induction t) auto
+
+lemma rbt_height_le: assumes "xrbt t" shows "height t \<le> 2 * log 2 (size1 t)"
+proof -
+ have "2 powr (height t / 2) \<le> 2 powr bheight t"
+ using rbt_height_bheight[OF assms] by (simp)
+ also have "\<dots> \<le> size1 t" using assms
+ by (simp add: powr_realpow bheight_size_bound xrbt_def)
+ finally have "2 powr (height t / 2) \<le> size1 t" .
+ hence "height t / 2 \<le> log 2 (size1 t)"
+ by(simp add: le_log_iff size1_def del: Int.divide_le_eq_numeral1(1))
+ thus ?thesis by simp
+
+ (* Todo: t_insert etc. using t_f_constant *)
+qed
+
+end
+
+print_locale xRBT_ins_del
+
+section\<open>Example 1: Intervals\<close>
+text\<open>See e.g. Cormen, Leiserson et al. \<comment> Introduction to Algorithms chapter 14.3\<close>
+
+type_synonym interval = "nat*nat"
+interpretation "nat\<times>nat" : linorder
+ oops \<comment>\<open>I failed -.-\<close>
+
+section\<open>Example 2: size information\<close>
+definition size :: "nat \<Rightarrow> 't \<Rightarrow> nat \<Rightarrow> nat" where
+ "size s1 _ s2 = Suc (s1 + s2)"
+
+definition t_size :: "nat \<Rightarrow> 't \<Rightarrow> nat \<Rightarrow> nat" where
+ "t_size s1 _ s2 = 1"--\<open>Assume constant, numbers stay low.\<close>
+
+interpretation xsize: xRBT_ins_del 0 t_size size
+ apply unfold_locales
+ unfolding t_size_def
+ apply auto
+ done
+
+fun select :: "('a,nat) xrbt \<Rightarrow> nat \<Rightarrow> 'a::linorder" where
+ "select (Node (_,x) l a r) i = (let sl = xsize.get_x l in if i=sl then a else if i < sl then select l i else select r (i - sl-1))"
+
+lemma inorder_size: "xsize.invx t \<Longrightarrow> length (inorder t) = xsize.get_x t"
+ apply (induction t)
+ by (auto simp: size_def)
+
+lemma "xsize.invx t \<Longrightarrow> i < xsize.get_x t \<Longrightarrow> inorder t ! i = select t i"
+ apply (induction t arbitrary: i)
+ apply (auto simp: nth_append Let_def size_def inorder_size)
+ done
+
+fun rank :: "('a,nat) xrbt \<Rightarrow> 'a::linorder \<Rightarrow> nat" where
+ "rank (Node (_,x) l a r) e = (if e = a then xsize.get_x l else if e < a then rank l e else xsize.get_x l + 1 + rank r e)"
+
+lemma sorted_append: "sorted (l1@l2) \<Longrightarrow> sorted l1 \<and> sorted l2"
+ apply (induction l1 rule: sorted.induct)
+ apply (auto simp: sorted_cons)
+ done
+
+lemma length_ok: "xsize.invx t \<Longrightarrow> sorted (inorder t) \<Longrightarrow> isin t e \<Longrightarrow> rank t e < length (inorder t)"
+ apply (induction t e rule: rank.induct)
+ apply auto
+ apply (meson less_SucI sorted_append trans_less_add1)
+ apply (simp add: inorder_size)
+ by (metis add_mono_thms_linordered_field(2) inorder_size sorted_append sorted_cons)
+
+lemma "xsize.invx t \<Longrightarrow> sorted (inorder t) \<Longrightarrow> isin t e \<Longrightarrow> inorder t !(rank t e) = e"
+ proof (induction t e rule: rank.induct)
+ case (1 uu x l a r e)
+ consider "e<a" | "e=a" | "e>a"
+ using antisym_conv3 by blast
+ then show ?case
+ using 1 apply cases
+ apply auto
+ apply (metis length_ok sorted_append nth_append)
+ apply (metis inorder_size nth_append_length)
+ apply (metis Suc_eq_plus1 ab_semigroup_add_class.add_ac(1)
+ inorder_size sorted_append nth_Cons_Suc nth_append_length_plus sorted_cons)
+ done
+ qed simp
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Hellauer_Fabian_f.hellauer@tum.de_656/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,2 @@
+Using temporary directory '/tmp/tmp.GoDGfWZVtX'
+Exercise path '/home/admin/suites/exbeoriginal' not found
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Hu_Shuwei_shuwei.hu@tum.de_651/Homework_Fibonacci.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,300 @@
+(* Author: Shuwei Hu *)
+
+theory Homework_Fibonacci
+ imports "~~/src/HOL/Library/Multiset"
+begin
+
+text \<open>
+This is an unsuccessful :( attempt of formalizing Fibonacci Heap.
+Here is three main difficulties and my possible solutions for them:
+ 1. @{term decrease_key} operation of any heap doesn't make much sense
+ in a functional setting, because it requires to identify a node, which
+ is trivial for imperative languages but requires an explicit "id" field
+ in functional languages. But @{term decrease_key} is the only reason why
+ we need a "mark" in a Fibonacci Heap. Here I abandoned the mark, which makes
+ the "fake Fibonacci Heap" effectively a collection of binomial heaps.
+
+ 2. The constant time of insertion and merging requires doubly linked list, which
+ is not even easy to just implement (possible in Haskell, but still not easy)
+ let alone verifying it. It is possible to just use list here and claiming that
+ there could be an efficient implementation in a practical functional language,
+ but after some brainstorm, I think there may be some potential confliction between
+ the nice properties of doubly linked list and the "functionality".
+ So I decided to implement a "mergeable list" by myself. It is a binary tree where
+ all values are stored in the leaves. It is functional. It supports constant time
+ merging by creating a new tree. However, it doesn't support deletion. This is why
+ I didn't implement del_min for the heap.
+
+ 3. The @{term consolidate} operation requires randomly accessible array. I used
+ @{type map} here as an alternative.
+\<close>
+
+section \<open>Mergeable Sequence\<close>
+datatype 'a nonempty_tree = Leaf 'a | Node "'a nonempty_tree" "'a nonempty_tree"
+term 0 (**)
+datatype 'a seq = EmptySeq ("[[]]") | TreeSeq "'a nonempty_tree"
+term set_seq
+term 0 (**)
+
+lemma set_nonempty_tree_not_empty[simp]: "set_nonempty_tree x \<noteq> {}"
+ by (induction x) auto
+
+lemma set_nonempty_tree_finite[simp]: "finite (set_nonempty_tree x)"
+ by (induction x) auto
+
+lemma set_seq_finite[simp]: "finite (set_seq x)"
+ by (cases x) auto
+term 0 (**)
+
+lemma seq_exhaust:
+ obtains "s = EmptySeq" | x where "s = TreeSeq (Leaf x)" | l r where "s = TreeSeq (Node l r)"
+ by (metis seq.exhaust nonempty_tree.exhaust)
+term 0 (**)
+
+lemma seq_induct:
+ assumes "P EmptySeq"
+ and "\<And>x. P (TreeSeq (Leaf x))"
+ and "\<And>l r. \<lbrakk>P (TreeSeq l); P (TreeSeq r)\<rbrakk> \<Longrightarrow> P (TreeSeq (Node l r))"
+ shows "P t"
+ apply (induction t)
+ subgoal by (simp only: assms)
+ subgoal for x by (induction x; simp only: assms)
+ done
+term 0 (**)
+
+fun append_seq :: "'a seq \<Rightarrow> 'a seq \<Rightarrow> 'a seq" (infixr "@@" 65) where
+ "append_seq EmptySeq r = r"
+| "append_seq l EmptySeq = l"
+| "append_seq (TreeSeq l) (TreeSeq r) = TreeSeq (Node l r)"
+
+fun SingleSeq :: "'a \<Rightarrow> 'a seq" ("[[_]]") where
+ "SingleSeq x = TreeSeq (Leaf x)"
+
+fun ConsSeq :: "'a \<Rightarrow> 'a seq \<Rightarrow> 'a seq" (infixr "##" 65) where
+ "ConsSeq x xt = append_seq (SingleSeq x) xt"
+term 0 (**)
+
+ (*
+fun foldr_nonempty_tree :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a nonempty_tree \<Rightarrow> 'b \<Rightarrow> 'b" where
+ "foldr_nonempty_tree f (Leaf x) = f x"
+| "foldr_nonempty_tree f (Node l r) = foldr_nonempty_tree f l \<circ> foldr_nonempty_tree f r"
+term 0 (**)
+*)
+
+fun flatten_seq :: "'a seq \<Rightarrow> 'a list" where
+ "flatten_seq EmptySeq = []"
+| "flatten_seq (TreeSeq (Leaf x)) = [x]"
+| "flatten_seq (TreeSeq (Node l r)) = flatten_seq (TreeSeq l) @ flatten_seq (TreeSeq r)"
+
+lemma flatten_seq_append[simp]:
+ "flatten_seq (l@@r) = flatten_seq l @ flatten_seq r"
+ by (cases l rule: seq_exhaust; cases r rule: seq_exhaust) auto
+term 0 (**)
+
+lemma set_seq_flatten[simp]:
+ "set_seq x = set (flatten_seq x)"
+ by (induction x rule: flatten_seq.induct) auto
+
+lemma set_seq_append_union: "set_seq (a@@b) = set_seq a \<union> set_seq b"
+ by auto
+term 0 (**)
+
+lemma flatten_map[simp]:
+ "flatten_seq (map_seq f xt) = map f (flatten_seq xt)"
+ by (induction xt rule: flatten_seq.induct) auto
+
+lemma flatten_not_Nil[simp]:
+ "xt \<noteq> [[]] \<Longrightarrow> flatten_seq xt \<noteq> []"
+ by (induction xt rule: flatten_seq.induct) auto
+term 0 (**)
+
+fun mset_seq :: "'a seq \<Rightarrow> 'a multiset" where
+ "mset_seq EmptySeq = {#}"
+| "mset_seq (TreeSeq (Leaf x)) = {#x#}"
+| "mset_seq (TreeSeq (Node l r)) = mset_seq (TreeSeq l) + mset_seq (TreeSeq r)"
+
+lemma mset_seq_flatten[simp]:
+ "mset_seq xt = mset (flatten_seq xt)"
+ by (induction xt rule: flatten_seq.induct) auto
+
+lemma (*append_seq_mest:*)
+ "mset_seq (l@@r) = mset_seq l + mset_seq r"
+ by auto
+
+section \<open>Fibonacci Heap\<close>
+datatype 'a hp = Hp (root: 'a) "'a hp seq"
+datatype 'a heap = Heap 'a "'a hp seq"
+term 0 (**)
+
+fun set_heap :: "'a heap \<Rightarrow> 'a set" where
+ "set_heap (Heap _ hps) = (\<Union>hp\<in>set_seq hps. set_hp hp)"
+
+term 0 (**)
+lemma [simp]: "set_hp hp \<noteq> {}"
+ by (cases hp) simp
+
+lemma [simp]: "hps \<noteq> [[]] \<Longrightarrow> set_heap (Heap m hps) \<noteq> {}"
+ by (cases hps) auto
+term 0 (**)
+
+lemma [simp]: "finite (set_hp hp)"
+ by (induction hp) auto
+term 0 (**)
+
+lemma [simp]: "finite (set_heap h)"
+ by (induction h) auto
+term 0 (**)
+
+definition "size_hp' \<equiv> size_hp (\<lambda>_. 0)"
+
+lemma measure_hp[measure_function]: "is_measure size_hp'" ..
+
+lemma measure_hp_relation[simp]: "hp \<in> set_seq hps \<Longrightarrow> size_hp' hp < size_hp' (Hp x hps)"
+ unfolding size_hp'_def by (induction hps rule: seq_induct) auto
+term 0 (**)
+
+fun mset_hp :: "'a hp \<Rightarrow> 'a multiset" where
+ "mset_hp (Hp x hps) = {#x#} + Union_mset(mset_seq(map_seq mset_hp hps))"
+term 0 (**)
+
+fun mset_heap :: "'a heap \<Rightarrow> 'a multiset" where
+ "mset_heap (Heap _ hps) = Union_mset(mset_seq(map_seq mset_hp hps))"
+term 0 (**)
+
+fun degree :: "'a hp \<Rightarrow> nat" where
+ "degree (Hp _ hps) = size hps"
+
+fun php :: "('a :: linorder) hp \<Rightarrow> bool" where
+ "php (Hp x hps) \<longleftrightarrow>
+ (\<forall>hp\<in>set_seq hps. php hp \<and> x \<le> root hp \<and> degree hp < size hps)
+ \<and> distinct (map degree (flatten_seq hps))"
+term 0 (**)
+
+fun pheap :: "('a :: linorder) heap \<Rightarrow> bool" where
+ "pheap (Heap x hps) \<longleftrightarrow> hps = [[]] \<or> ((\<forall>hp\<in>set_seq hps. php hp) \<and> x = Min (set_seq (map_seq root hps)))"
+term 0 (**)
+
+lemma set_hp_alt: "set_hp (Hp x hps) = {x} \<union> (\<Union>hp\<in>set_seq hps. set_hp hp)"
+ by simp
+term 0 (**)
+
+lemma hp_root_min: "php hp \<Longrightarrow> root hp = Min (set_hp hp)"
+ by (induction hp) (fastforce intro!: Min_eqI[symmetric])
+term 0 (**)
+
+definition empty_heap :: "'a heap" where
+ "empty_heap \<equiv> Heap undefined [[]]"
+
+fun is_empty :: "'a heap \<Rightarrow> bool" where
+ "is_empty (Heap _ hps) \<longleftrightarrow> hps = [[]]"
+term 0 (**)
+
+lemma pheap_empty[simp]: "pheap empty_heap"
+ unfolding empty_heap_def by simp
+term 0 (**)
+
+lemma empty_is_empty[simp]: "is_empty empty_heap"
+ unfolding empty_heap_def by simp
+term 0 (**)
+
+fun merge :: "('a :: linorder) heap \<Rightarrow> 'a heap \<Rightarrow> 'a heap" where
+ "merge (Heap m0 hps0) (Heap m1 hps1) =
+ Heap (if hps0=[[]] then m1 else if hps1=[[]] then m0 else min m0 m1) (hps0@@hps1)"
+term 0 (**)
+
+lemma pheap_merge: "\<lbrakk>pheap h0; pheap h1\<rbrakk> \<Longrightarrow> pheap (merge h0 h1)"
+ by (induction h0 h1 rule: merge.induct) (auto simp: Min_Un)
+term 0 (**)
+
+lemma mset_merge: "mset_heap (merge h0 h1) = mset_heap h0 + mset_heap h1"
+ by (induction h0 h1 rule: merge.induct) auto
+term 0 (**)
+
+fun get_min :: "('a :: linorder heap) \<Rightarrow> 'a" where
+ "get_min (Heap m hps) = m"
+
+lemma get_min_in: "\<lbrakk>pheap h; \<not>is_empty h\<rbrakk> \<Longrightarrow> get_min h \<in> set_heap h"
+proof (induction h rule: get_min.induct)
+ case (1 m hps)
+ hence "\<exists>hp\<in>set_seq hps. m = root hp"
+ by (auto intro: bex_imageD[of root])
+ thus ?case
+ by (auto intro: hp.set_sel)
+qed
+
+lemma get_min_min: "\<lbrakk>pheap h; \<not>is_empty h; x\<in>set_heap h\<rbrakk> \<Longrightarrow> get_min h \<le> x"
+proof (induction h rule: get_min.induct)
+ case (1 m hps)
+ then obtain hp where *:"hp\<in>set_seq hps \<and> x\<in>set_hp hp"
+ by auto
+ from 1 * have "get_min (Heap m hps) \<le> root hp"
+ by auto
+ also from 1 * have "\<dots> \<le> x"
+ by (auto simp: hp_root_min)
+ finally show ?case .
+qed
+
+fun make_heap :: "('a :: linorder) \<Rightarrow> 'a heap" where
+ "make_heap x = Heap x (SingleSeq (Hp x [[]]))"
+term Nil
+
+lemma pheap_make_heap: "pheap (make_heap x)"
+ by simp
+term 0 (**)
+
+fun insert :: "('a :: linorder) \<Rightarrow> 'a heap \<Rightarrow> 'a heap" where
+ "insert x h = merge h (make_heap x)"
+
+lemma pheap_insert: "pheap h \<Longrightarrow> pheap (insert x h)"
+ by (auto intro: pheap_merge pheap_make_heap)
+
+lemma mset_insert: "pheap h \<Longrightarrow> mset_heap (insert x h) = {#x#} + mset_heap h"
+ by (auto simp: mset_merge)
+
+fun meld :: "('a :: linorder) hp \<Rightarrow> 'a hp \<Rightarrow> 'a hp" where
+ "meld (Hp x0 hps0) (Hp x1 hps1) =
+ (if x0\<le>x1 then Hp x0 (Hp x1 hps1 ## hps0) else (Hp x1 (Hp x0 hps0 ## hps1)))"
+
+function consolidate_step :: "nat \<Rightarrow> nat \<Rightarrow> ('a :: linorder) hp \<Rightarrow> (nat \<rightharpoonup> 'a hp) \<Rightarrow> (nat \<rightharpoonup> 'a hp)" where
+ "consolidate_step max_dim d hp t =
+ (if d < max_dim
+ then (case t d of
+ None \<Rightarrow> t(d\<mapsto>hp)
+ | Some hp' \<Rightarrow> consolidate_step max_dim (Suc d) (meld hp hp') (t(d:=None)))
+ else t)"
+ by pat_completeness auto
+termination
+ by (relation "measure (\<lambda>(d,i,_,_). d-i)") auto
+
+fun consolidate :: "nat \<Rightarrow> ('a :: linorder) hp seq \<Rightarrow> (nat \<rightharpoonup> 'a hp) \<Rightarrow> (nat \<rightharpoonup> 'a hp)" where
+ "consolidate max_dim EmptySeq t = t"
+| "consolidate max_dim (TreeSeq (Leaf x)) t =
+ consolidate_step max_dim (degree x) x t"
+| "consolidate max_dim (TreeSeq (Node l r)) t =
+ consolidate max_dim (TreeSeq l) (consolidate max_dim (TreeSeq r) t)"
+
+lemma meld_php:
+ "\<lbrakk>php hp0; php hp1; degree hp0 = degree hp1\<rbrakk> \<Longrightarrow> php (meld hp0 hp1)"
+proof (induction hp0 hp1 rule: meld.induct)
+ case (1 x0 hps0 x1 hps1)
+ have sizeeq: "size hps0 = size hps1"
+ using \<open>degree (Hp x0 hps0) = degree (Hp x1 hps1)\<close> by auto
+ { assume le: "x0 \<le> x1"
+ then have "meld (Hp x0 hps0) (Hp x1 hps1) = Hp x0 (Hp x1 hps1 ## hps0)"
+ by simp
+ have "php (Hp x0 (Hp x1 hps1 ## hps0))"
+ proof -
+ fix hp assume "hp \<in> set_seq (Hp x1 hps1 ## hps0)"
+ then have *: "hp = Hp x1 hps1 \<or> hp \<in> set_seq hps0"
+ by auto
+ have "php hp"
+ using * \<open>php (Hp x0 hps0)\<close> \<open>php (Hp x1 hps1)\<close> by auto
+ moreover
+ have "x0 \<le> root hp"
+ using * \<open>php (Hp x0 hps0)\<close> \<open>php (Hp x1 hps1)\<close> le by auto
+ moreover
+ have "size (Hp x1 hps1 ## hps0) = 1 + size hps0"
+ oops
+end
+
+
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Hu_Shuwei_shuwei.hu@tum.de_651/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,2 @@
+Using temporary directory '/tmp/tmp.RgdAPY2D56'
+Exercise path '/home/admin/suites/exbeoriginal' not found
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Hutzler_Matthias_ga95luy@mytum.de_638/digits_arithmetic.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,475 @@
+(*\<Rightarrow> Tobias*)
+(* Author: Matthias Hutzler
+
+ Some arbitrary precision arithmetic with lists of binary digits.
+
+ datatype number = Number sign exponent "digit list"
+ re :: "number \<Rightarrow> real"
+
+ Actually, not much arithmetic at all. (Didn't get to it.)
+ Instead, there is a normal form for digit representations.
+*)
+
+theory digits_arithmetic
+ imports Complex_Main
+begin
+
+
+section "Digit representations and the reals they represent"
+
+datatype digit = Zero | One
+
+datatype sign = Plus | Minus
+
+type_synonym exponent = nat
+
+ (* Type of numbers represented by digit lists. *)
+datatype number = Number sign exponent "digit list"
+
+fun fromDigit :: "digit \<Rightarrow> real" where
+ "fromDigit Zero = 0" |
+ "fromDigit One = 1"
+
+ (* Turn a list of binary digits starting just after the decimal point
+ into a number between 0 and 1 *)
+fun re_digits :: "digit list \<Rightarrow> real" where
+ "re_digits [] = 0" |
+ "re_digits (d#ds) = 0.5 * (fromDigit d + re_digits ds)"
+
+lemma fromDigit_lower_bound[simp] : "0 \<le> fromDigit d"
+ by (cases d) auto
+lemma fromDigit_upper_bound[simp] : "fromDigit d \<le> 1"
+ by (cases d) auto
+
+lemma re_digits_lower_bound[simp] : "0 \<le> re_digits ds"
+ by (induction ds) auto
+
+lemma re_digits_upper_bound[simp] : "re_digits ds < 1"
+proof (induction ds)
+ case Nil
+ show ?case by simp
+next
+ case (Cons a ds)
+ have "2 * re_digits (a # ds) \<le> 1 + re_digits ds" by auto
+ also have "... < 2" using Cons by auto
+ finally show ?case by simp
+qed
+
+lemma [simp]: "re_digits (ds@[Zero]) = re_digits ds"
+ by (induction ds rule: re_digits.induct) auto
+
+fun fromSign :: "sign \<Rightarrow> real" where
+ "fromSign Plus = 1" |
+ "fromSign Minus = -1"
+
+lemma [simp]: "fromSign s \<noteq> 0"
+ by (cases s) auto
+
+fun fromExponent :: "exponent \<Rightarrow> real" where
+ "fromExponent e = 2^e"
+
+lemma [simp]: "fromExponent e > 0"
+ by auto
+
+ (* Realization of digit representations as real numbers. *)
+definition re :: "number \<Rightarrow> real" where
+ "re n = (case n of (Number s e ds) \<Rightarrow> fromSign s * fromExponent e * re_digits ds)"
+
+lemma re_not_inj_aux: "re (Number Plus 0 [One]) = re (Number Plus 1 [Zero, One])"
+ unfolding re_def by simp
+
+lemma re_not_inj: "\<not> inj re"
+ unfolding inj_on_def
+ using re_not_inj_aux by auto
+
+lemma re_not_surj: "\<not> surj re"
+ sorry (* 1/3 is a counterexample *)
+
+
+section "Normal form for digit representations"
+
+fun no_trailing_zeros :: "number \<Rightarrow> bool" where
+ "no_trailing_zeros (Number s e ds) \<longleftrightarrow> \<not>(\<exists>ds'. ds=ds'@[Zero])"
+
+fun minimal_exponent :: "number \<Rightarrow> bool" where
+ "minimal_exponent (Number s e ds) \<longleftrightarrow> e=0 \<or> (\<exists>ds'. ds=One#ds')"
+
+fun plus_if_empty_digits :: "number \<Rightarrow> bool" where
+ "plus_if_empty_digits (Number s e ds) \<longleftrightarrow> (ds=[] \<longrightarrow> s = Plus)"
+
+ (* normal form for digit representations *)
+definition normal :: "number \<Rightarrow> bool" where
+ "normal n \<equiv> no_trailing_zeros n \<and> minimal_exponent n \<and> plus_if_empty_digits n"
+
+function (sequential) reduce_exponent :: "number \<Rightarrow> number" where
+ "reduce_exponent (Number s (Suc e) (Zero#ds)) = reduce_exponent (Number s e ds)" |
+ "reduce_exponent (Number s e []) = Number s 0 []" |
+ "reduce_exponent (Number s e ds) = Number s e ds"
+ by pat_completeness auto
+termination
+ by (relation "measure (\<lambda>n. case n of (Number s e ds) \<Rightarrow> e)") auto
+
+lemma reduce_exponent[simp]: "minimal_exponent (reduce_exponent n)"
+ by (induction n rule: reduce_exponent.induct) auto
+
+lemma re_reduce_exponent[simp]: "re (reduce_exponent n) = re n"
+ unfolding re_def
+ by (induction n rule: reduce_exponent.induct) auto
+
+fun dropZeros :: "digit list \<Rightarrow> digit list" where
+ "dropZeros (Zero#ds) = dropZeros ds" |
+ "dropZeros ds = ds"
+
+lemma [simp]: "\<not> rev (dropZeros l) = l'@[Zero]"
+ by (induction l rule: dropZeros.induct) auto
+
+lemma [simp]: "dropZeros (l1@[One]) = dropZeros l1 @[One]"
+ by (induction l1 rule: dropZeros.induct) auto
+
+lemma [simp]: "\<exists>l'. l=One#l' \<Longrightarrow> \<exists>l'. (rev (dropZeros (rev l))) = One#l'"
+ by (induction l rule: dropZeros.induct) auto
+
+lemma [simp]: "re_digits (rev (dropZeros rds)) = re_digits (rev rds)"
+ by (induction rds rule: dropZeros.induct) auto
+
+fun cut_trailing_zeros :: "number \<Rightarrow> number" where
+ "cut_trailing_zeros (Number s e ds) = Number s e (rev (dropZeros (rev ds)))"
+
+lemma cut_trailing_zeros[simp]: "no_trailing_zeros (cut_trailing_zeros n)"
+ by (induction n rule: cut_trailing_zeros.induct) auto
+
+lemma minimal_exponent_cut_trailing_zeros_aux1:
+ "minimal_exponent (cut_trailing_zeros (Number s 0 ds))"
+ by auto
+lemma minimal_exponent_cut_trailing_zeros_aux2:
+ "minimal_exponent (cut_trailing_zeros (Number s e (One#ds)))"
+ by auto
+
+lemma minimal_exponent_cut_trailing_zeros[simp]:
+ "minimal_exponent n \<Longrightarrow> minimal_exponent (cut_trailing_zeros n)"
+ by (metis minimal_exponent.elims(2)
+ minimal_exponent_cut_trailing_zeros_aux1 minimal_exponent_cut_trailing_zeros_aux2)
+
+lemma re_cut_trailing_zeros[simp]: "re (cut_trailing_zeros n) = re n"
+ unfolding re_def
+ by (induction n rule: cut_trailing_zeros.induct) auto
+
+fun check_empty_digits :: "number \<Rightarrow> number" where
+ "check_empty_digits (Number s e []) = Number Plus 0 []" |
+ "check_empty_digits n = n"
+
+lemma check_empty_digits[simp]: "plus_if_empty_digits (check_empty_digits n)"
+ by (induction n rule: check_empty_digits.induct) auto
+
+lemma minimal_exponent_check_empty_digits[simp]:
+ "minimal_exponent n \<Longrightarrow> minimal_exponent (check_empty_digits n)"
+ by (induction n rule: check_empty_digits.induct) auto
+
+lemma no_trailing_zeros_check_empty_digits[simp]:
+ "no_trailing_zeros n \<Longrightarrow> no_trailing_zeros (check_empty_digits n)"
+ by (induction n rule: check_empty_digits.induct) auto
+
+lemma [simp]: "re (Number s e []) = 0"
+ unfolding re_def
+ by simp
+
+lemma re_check_empty_digits[simp]: "re (check_empty_digits n) = re n"
+ by (induction n rule: check_empty_digits.induct) auto
+
+definition normal_form :: "number \<Rightarrow> number" where
+ "normal_form \<equiv> check_empty_digits \<circ> cut_trailing_zeros \<circ> reduce_exponent"
+
+lemma normal_normal_form: "normal (normal_form n)"
+ unfolding normal_def normal_form_def
+ by (auto)
+
+lemma re_normal_form: "re (normal_form n) = re n"
+ unfolding normal_form_def
+ by auto
+
+
+section "Injectivity on representations in normal form"
+
+lemma sum_of_nonnegative_0: "a\<ge>0 \<Longrightarrow> b\<ge>0 \<Longrightarrow> (a::real) + b = 0 \<Longrightarrow> a=0 \<and> b=0"
+ by simp
+
+lemma re_digits_0: "re_digits ds = 0 \<Longrightarrow> \<forall>d\<in>set ds. d=Zero"
+proof (induction ds rule: re_digits.induct)
+ case 1
+ show ?case by simp
+next
+ case (2 d ds)
+ hence "fromDigit d + re_digits ds = 0" by simp
+ with sum_of_nonnegative_0[OF fromDigit_lower_bound[of d] re_digits_lower_bound[of ds]]
+ have d0: "fromDigit d = 0"
+ and ds0: "re_digits ds = 0"
+ by auto
+ from d0 have "d = Zero" by (cases d) auto
+ with ds0 2(1) show ?case by auto
+qed
+
+lemma re_0_only_zeros: "re (Number s e ds) = 0 \<Longrightarrow> \<forall>d\<in>set ds. d=Zero"
+ unfolding re_def
+ using re_digits_0 by auto
+
+lemma [simp]: "normal (Number s e []) \<longleftrightarrow> s=Plus \<and> e=0"
+ unfolding normal_def
+ by auto
+
+lemma no_trailing_zeros_only_zeros:
+ "\<not>(\<exists>ds'. ds=ds'@[Zero]) \<Longrightarrow> \<forall>d\<in>set ds. d=Zero \<Longrightarrow> ds = []"
+ by (metis append_butlast_last_id last_in_set)
+
+lemma re_0:
+ assumes "normal n"
+ assumes "re n = 0"
+ shows "n = Number Plus 0 []"
+proof (cases n)
+ case (Number s e ds)
+ with assms(2) re_0_only_zeros have "\<forall>d\<in>set ds. d=Zero" by simp
+ with no_trailing_zeros_only_zeros assms(1) normal_def Number have "ds=[]" by simp
+ from this Number assms(1) show ?thesis by auto
+qed
+
+lemma [simp]: "0 \<le> re (Number Plus e ds)"
+ unfolding re_def by simp
+
+lemma [simp]: "re (Number Minus e ds) \<le> 0"
+ unfolding re_def by simp
+
+ (* helper function to get a nice induction rule *)
+fun digits_equal :: "digit list \<Rightarrow> digit list \<Rightarrow> bool" where
+ "digits_equal [] [] \<longleftrightarrow> True" |
+ "digits_equal (c#cs) (d#ds) \<longleftrightarrow> c=d \<and> digits_equal cs ds" |
+ "digits_equal _ _ \<longleftrightarrow> False"
+
+lemma re_digits_eq_head:
+ assumes "re_digits (c#cs) = re_digits (d#ds)"
+ shows "c=d"
+proof (cases c)
+ case Zero
+ with assms have "fromDigit d + re_digits ds = re_digits cs" by simp
+ hence "fromDigit d + re_digits ds < 1" by auto
+ with re_digits_lower_bound[of ds] have "fromDigit d < 1" by linarith
+ with Zero show ?thesis by (cases d) auto
+next
+ case One
+ with assms have "fromDigit d + re_digits ds = 1 + re_digits cs" by simp
+ hence "fromDigit d + re_digits ds \<ge> 1" by auto
+ with re_digits_upper_bound[of ds] have "fromDigit d > 0" by linarith
+ with One show ?thesis by (cases d) auto
+qed
+
+lemma re_digits_eq_tail:
+ assumes "re_digits (c#cs) = re_digits (d#ds)"
+ shows "re_digits cs = re_digits ds"
+proof -
+ from assms re_digits_eq_head have "c=d" by simp
+ thus ?thesis using assms by auto
+qed
+
+lemma last_of_tail_not_zero: "\<not>(\<exists>l'. b#l = l'@[Zero]) \<Longrightarrow> \<not>(\<exists>l'. l = l'@[Zero])"
+ by auto
+
+lemma re_digits_inj: "\<not>(\<exists>ds'. ds1=ds'@[Zero]) \<Longrightarrow> \<not>(\<exists>ds'. ds2=ds'@[Zero]) \<Longrightarrow>
+ re_digits ds1 = re_digits ds2 \<Longrightarrow> ds1=ds2"
+proof (induction ds1 ds2 rule: digits_equal.induct)
+ case 1
+ show ?case by simp
+next
+ case (2 c cs d ds)
+ with re_digits_eq_head have "c = d" by simp
+ moreover from 2(2) last_of_tail_not_zero have "\<not>(\<exists>ds'. cs=ds'@[Zero])" by auto
+ moreover from 2(3) last_of_tail_not_zero have "\<not>(\<exists>ds'. ds=ds'@[Zero])" by auto
+ ultimately show ?case
+ using 2(1) re_digits_eq_tail[OF 2(4)] by auto
+next
+ case ("3_1" v va)
+ let ?ds = "butlast (v#va)"
+ let ?d = "last (v#va)"
+ have "?ds@[?d] = v#va" by auto
+ with "3_1"(1) have dOne: "?d=One" by (cases ?d) metis
+ from "3_1"(3) re_digits_0[of "v#va"] have "\<forall>d\<in>set (v#va). d=Zero" by simp
+ with dOne last_in_set[of "v#va"] have False by auto
+ thus ?case by simp
+next
+ case ("3_2" v va)
+ let ?ds = "butlast (v#va)"
+ let ?d = "last (v#va)"
+ have "?ds@[?d] = v#va" by auto
+ with "3_2"(2) have dOne: "?d=One" by (cases ?d) metis
+ from "3_2"(3) re_digits_0[of "v#va"] have "\<forall>d\<in>set (v#va). d=Zero" by simp
+ with dOne last_in_set[of "v#va"] have False by auto
+ thus ?case by simp
+qed
+
+lemma abs_re: "abs (re (Number s e ds)) = 2^e * re_digits ds"
+ unfolding re_def
+ by (cases s) auto
+
+lemma abs_re_upper_bound: "abs (re (Number s e ds)) < 2^e"
+ using abs_re by auto
+
+lemma abs_re_lower_bound:
+ "normal (Number s e ds) \<Longrightarrow> e\<noteq>0 \<Longrightarrow> 2^e \<le> 2 * abs (re (Number s e ds))"
+ unfolding normal_def
+ using abs_re by auto
+
+lemma re_inj_aux:
+ assumes "normal n1"
+ assumes "normal n2"
+ assumes "re n1 = re n2"
+ shows "n1 = n2"
+proof (cases n1)
+ case (Number s1 e1 ds1)
+ note n1 = this
+ let ?x = "re n1"
+ show ?thesis
+ proof (cases n2)
+ case (Number s2 e2 ds2)
+ note n2 = this
+ show ?thesis
+ proof cases
+ assume "?x = 0"
+ then show ?thesis
+ using re_0[of n1] re_0[of n2] assms by auto
+ next
+ assume "?x \<noteq> 0"
+
+ have signs: "s1 = s2"
+ proof (cases s1)
+ case Plus
+ with n1 have "?x \<ge> 0" by auto
+ with `?x\<noteq>0` assms(3) have "\<not>(re n2 \<le> 0)" by auto
+ then show ?thesis using n2 Plus by (cases s2) auto
+ next
+ case Minus
+ with n1 have "?x \<le> 0" by auto
+ with `?x\<noteq>0` assms(3) have "\<not>(0 \<le> re n2)" by auto
+ then show ?thesis using n2 Minus by (cases s2) auto
+ qed
+
+ have "\<not>(e1<e2)"
+ proof
+ assume "e1 < e2"
+ hence e1e2: "e1+1 \<le> e2" by simp
+ hence e1e2': "(2::real)^(e1+1) \<le> 2^e2" by (rule power_increasing) simp
+ from e1e2 have "e2\<noteq>0" by simp
+ have "2 * abs ?x < 2^(e1+1)" using abs_re_upper_bound n1 by auto
+ also have "... \<le> 2^e2" using e1e2' .
+ also have "... \<le> 2 * abs ?x"
+ using abs_re_lower_bound n2 assms(2) assms(3) `e2\<noteq>0` by auto
+ finally show False by simp
+ qed
+ moreover have "\<not>(e2<e1)"
+ proof
+ assume "e2 < e1"
+ hence e1e2: "e2+1 \<le> e1" by simp
+ hence e1e2': "(2::real)^(e2+1) \<le> 2^e1" by (rule power_increasing) simp
+ from e1e2 have "e1\<noteq>0" by simp
+ have "2 * abs ?x < 2^(e2+1)" using abs_re_upper_bound n2 assms(3) by auto
+ also have "... \<le> 2^e1" using e1e2' .
+ also have "... \<le> 2 * abs ?x"
+ using abs_re_lower_bound n1 assms(1) `e1\<noteq>0` by auto
+ finally show False by simp
+ qed
+ ultimately have exponents: "e1 = e2" by simp
+
+ have "abs (re n1) = abs (re n2)" using assms(3) by simp
+ hence "2^e1 * re_digits ds1 = 2^e2 * re_digits ds2" using n1 n2 abs_re by simp
+ hence "re_digits ds1 = re_digits ds2" using exponents by simp
+ hence digits: "ds1 = ds2"
+ using re_digits_inj[of ds1 ds2] assms(1) assms(2) normal_def n1 n2
+ by auto
+
+ show ?thesis using signs exponents digits n1 n2 by simp
+ qed
+ qed
+qed
+
+definition normal_form_numbers :: "number set" where
+ "normal_form_numbers \<equiv> Collect normal"
+
+theorem re_inj_on_normals: "inj_on re normal_form_numbers"
+ unfolding inj_on_def normal_form_numbers_def
+ using re_inj_aux by auto
+
+
+section "Negation and addition"
+
+fun negate :: "number \<Rightarrow> number" where
+ "negate (Number Plus e ds) = Number Minus e ds" |
+ "negate (Number Minus e ds) = Number Plus e ds"
+
+lemma re_negate[simp] : "re (negate n) = - re n"
+ unfolding re_def
+ by (induction n rule: negate.induct) auto
+
+fun increaseExponent :: "nat \<Rightarrow> number \<Rightarrow> number" where
+ "increaseExponent 0 n = n" |
+ "increaseExponent (Suc a) (Number s e ds) = increaseExponent a (Number s (Suc e) (Zero#ds))"
+
+lemma re_increaseExponent[simp] : "re (increaseExponent a n) = re n"
+ unfolding re_def
+ by (induction a n rule: increaseExponent.induct) auto
+
+ (* The 'full adder' logic. *)
+fun add_three_digits :: "digit \<Rightarrow> digit \<Rightarrow> digit \<Rightarrow> digit * digit" where
+ "add_three_digits Zero Zero d = (Zero, d)" |
+ "add_three_digits One One d = (One, d)" |
+ "add_three_digits Zero One Zero = (Zero, One)" |
+ "add_three_digits Zero One One = (One, Zero)" |
+ "add_three_digits One Zero Zero = (Zero, One)" |
+ "add_three_digits One Zero One = (One, Zero)"
+
+lemma add_three_digits_correct : "(c, d) = add_three_digits d1 d2 d3 \<Longrightarrow>
+ 2 * fromDigit c + fromDigit d = fromDigit d1 + fromDigit d2 + fromDigit d3"
+ by (induction d1 d2 d3 rule: add_three_digits.induct) auto
+
+ (* Add two lists of binary digits, returning a carry and a new list of digits. *)
+fun add_digits :: "digit list \<Rightarrow> digit list \<Rightarrow> digit * digit list" where
+ "add_digits [] ds = (Zero, ds)" |
+ "add_digits cs [] = (Zero, cs)" |
+ "add_digits (c#cs) (d#ds) =
+ (let (carry0, ds') = add_digits cs ds in
+ (let (carry, d') = add_three_digits c d carry0 in
+ (carry, d'#ds') ))"
+
+lemma cancel_2_helper: "2*(a::real) = 2*b \<Longrightarrow> a = b" by simp
+
+lemma re_add_digits : "(carry, ds') = add_digits cs ds \<Longrightarrow>
+ fromDigit carry + re_digits ds' = re_digits cs + re_digits ds"
+proof (induction cs ds arbitrary: carry ds' rule: add_digits.induct)
+ case (1 ds)
+ thus ?case by simp
+next
+ case (2 v va)
+ thus ?case by simp
+next
+ case (3 c cs d ds)
+ let ?carry = "fst (add_digits cs ds)"
+ let ?ds' = "snd (add_digits cs ds)"
+ from 3(1) have IH: "fromDigit ?carry + re_digits ?ds' = re_digits cs + re_digits ds" by simp
+ from 3(2) have carry: "carry = fst (add_three_digits c d ?carry)" by (auto split: prod.splits(2))
+ let ?d' = "snd (add_three_digits c d ?carry)"
+ from 3(2) have "ds' = ?d' # ?ds'" by (simp split: prod.splits(2))
+ hence "2*(fromDigit carry + re_digits ds') = 2*fromDigit carry + fromDigit ?d' + re_digits ?ds'"
+ by simp
+ also have "... = fromDigit c + fromDigit d + fromDigit ?carry + re_digits ?ds'"
+ using carry add_three_digits_correct by simp
+ also have "... = fromDigit c + re_digits cs + fromDigit d + re_digits ds"
+ using IH by simp
+ also have "... = 2 * re_digits (c#cs) + fromDigit d + re_digits ds"
+ by simp
+ also have "... = 2 * (re_digits (c#cs) + re_digits (d#ds))"
+ by simp
+ finally show ?case by (rule cancel_2_helper)
+ (* Why does (rule iffD1[OF mult_left_cancel[of 2]]) not work here? *)
+qed
+
+fun add :: "number \<Rightarrow> number \<Rightarrow> number" where
+ "add n1 n2 = undefined"
+
+lemma re_add[simp]: "re (add n1 n2) = re n1 + re n2"
+ sorry
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Hutzler_Matthias_ga95luy@mytum.de_638/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,2 @@
+Using temporary directory '/tmp/tmp.Nvlq2O7GMq'
+Exercise path '/home/admin/suites/exbeoriginal' not found
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Keinholz_Jonas_jonas.keinholz@tum.de_658/project_euler.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,480 @@
+(*\<Rightarrow> Max *)
+
+section \<open>Formally verified solutions to Project Euler problems\<close>
+
+text \<open>
+ Formally verified solutions to the first two Project Euler problems (https://projecteuler.net)
+
+ Requirements:
+ * Find a solution that is efficient enough to get accepted by Hacker Rank
+ (https://www.hackerrank.com/contests/projecteuler/challenges)
+ * Code generation from the solution must be possible
+ * Prove correctness
+\<close>
+
+theory project_euler
+ imports
+ Main
+ "~~/src/HOL/Library/Code_Target_Numeral"
+begin
+
+subsection \<open>Project Euler #1: Multiples of 3 and 5\<close>
+
+text \<open>
+ Find the sum of all the multiples of 3 or 5 below $n$.
+\<close>
+
+definition solve001 :: "nat \<Rightarrow> nat" where
+ "solve001 n = \<Sum>{k \<in> {..<n}. k mod 3 = 0 \<or> k mod 5 = 0}"
+
+subsubsection \<open>Naive approach\<close>
+
+text \<open>
+ First, we find an executable version of @{term solve001}. This won't be efficient enough, though.
+\<close>
+
+definition solve001_naive :: "nat \<Rightarrow> nat" where
+ "solve001_naive n = sum_list [k\<leftarrow>[1..<n]. k mod 3 = 0 \<or> k mod 5 = 0]"
+
+lemma solve001_naive_correctness: "solve001_naive = solve001"
+proof (standard, goal_cases elem)
+ case (elem n)
+ have "solve001 n = sum_list [k\<leftarrow>[0..<n]. k mod 3 = 0 \<or> k mod 5 = 0]"
+ proof -
+ have "comm_monoid_list_set op + (0::nat)" using sum.comm_monoid_list_set_axioms by auto
+ then show ?thesis
+ unfolding solve001_def sum_def sum_list_def
+ using comm_monoid_list_set.distinct_set_conv_list[
+ where f = "op +" and
+ g = "\<lambda>k. k" and
+ xs = "[k\<leftarrow>[0..<n]. k mod 3 = 0 \<or> k mod 5 = 0]"]
+ by auto
+ qed
+ also have "\<dots> = solve001_naive n"
+ unfolding solve001_naive_def using upt_rec[where i = 0] by auto
+ finally show ?case by auto
+qed
+
+subsubsection \<open>Arithmetic approach\<close>
+
+text \<open>
+ We can compute the sum of multiples of $k$ below $n$ using an arithmetic progression. This
+ will be more efficient.
+\<close>
+
+abbreviation multiples :: "nat \<Rightarrow> nat \<Rightarrow> nat set" where
+ "multiples n m \<equiv> {k \<in> {..<n}. k mod m = 0}"
+
+lemma solve001_multiples: "solve001 n = \<Sum>(multiples n 3) + \<Sum>(multiples n 5) - \<Sum>(multiples n 15)"
+proof -
+ have "solve001 n = \<Sum>{k \<in> {..<n}. k mod 3 = 0 \<or> k mod 5 = 0}"
+ unfolding solve001_def by auto
+ also have "\<dots> = \<Sum>(multiples n 3) + \<Sum>(multiples n 5) - \<Sum>(multiples n 15)"
+ proof -
+ have "{k \<in> {..<n}. k mod 3 = 0 \<or> k mod 5 = 0} = multiples n 3 \<union> multiples n 5"
+ by auto
+ moreover have "multiples n 3 \<inter> multiples n 5 = multiples n 15"
+ proof -
+ have "multiples n 3 \<inter> multiples n 5 = {k \<in> {..<n}. k mod 3 = 0 \<and> k mod 5 = 0}"
+ by blast
+ also have "\<dots> = multiples n 15"
+ proof -
+ have "\<And>k :: nat. k mod 3 = 0 \<and> k mod 5 = 0 \<longleftrightarrow> k mod 15 = 0" by presburger
+ then show ?thesis by auto
+ qed
+ finally show ?thesis .
+ qed
+ ultimately show ?thesis using sum_Un_nat[where f = "\<lambda>k. k"] by auto
+ qed
+ finally show ?thesis .
+qed
+
+lemma Sum_f:
+ fixes f :: "nat \<Rightarrow> nat"
+ assumes "inj f"
+ shows "\<Sum>{f k | k. k \<le> n} = sum f {..n}"
+proof (induction n)
+ case (Suc n)
+ have "\<Sum>{f k | k. k \<le> Suc n} = f (Suc n) + \<Sum>({f k | k. k \<le> Suc n} - {f (Suc n)})"
+ unfolding sum_def
+ using sum.comm_monoid_set_axioms
+ by (rule comm_monoid_set.remove) auto
+ also have "\<dots> = f (Suc n) + \<Sum>{f k | k. k \<le> n}"
+ proof -
+ have "{f k | k. k \<le> Suc n} - {f (Suc n)} = {f k | k. k \<le> n}"
+ proof (standard, goal_cases LTR RTL)
+ case LTR
+ show ?case
+ proof (safe, goal_cases elem)
+ case (elem _ k)
+ moreover then have "k \<noteq> Suc n" by auto
+ ultimately show ?case by auto
+ qed
+ next
+ case RTL
+ show ?case
+ proof (safe, goal_cases)
+ case (2 x k)
+ then show ?case using assms injD by fastforce
+ qed auto
+ qed
+ then show ?thesis by auto
+ qed
+ also have "\<dots> = f (Suc n) + sum f {..n}" using Suc by auto
+ also have "\<dots> = sum f {..Suc n}" by auto
+ finally show ?case .
+qed auto
+
+text \<open>
+ Needed for some algebraic transformations.
+\<close>
+
+lemma Sum_multiples_even:
+ fixes n :: nat
+ shows "even (n * (n + 1))"
+proof -
+ have "n mod 2 = 0 \<or> (n + 1) mod 2 = 0" by presburger
+ then show ?thesis by auto
+qed
+
+lemma Sum_multiples:
+ fixes n m :: nat
+ shows "\<Sum>(multiples n m) = (m * ((n - 1) div m) * ((n - 1) div m + 1)) div 2"
+proof (cases "n = 0 \<or> m = 0")
+ case False
+ then have "n \<noteq> 0" "m \<noteq> 0" by auto
+
+ have "multiples n m = {m * q | q. q \<le> (n - 1) div m}"
+ proof (standard, goal_cases LTR RTL)
+ case LTR
+ show ?case
+ proof (standard, goal_cases elem)
+ case (elem k)
+ then have "k < n" "k mod m = 0" by auto
+ then obtain q where "k = m * q" by auto
+ moreover have "q \<le> (n - 1) div m"
+ proof -
+ have "q = k div m" using \<open>m \<noteq> 0\<close> \<open>k = m * q\<close> by auto
+ also have "\<dots> \<le> (n - 1) div m " using div_le_mono \<open>k < n\<close> by auto
+ finally show ?thesis .
+ qed
+ ultimately show ?case by auto
+ qed
+ next
+ case RTL
+ show ?case
+ proof (standard, goal_cases elem)
+ case (elem k)
+ then obtain q where q: "q \<le> (n - 1) div m" "k = m * q" by auto
+ moreover have "k < n"
+ proof -
+ have "k \<le> m * ((n - 1) div m)" using q by auto
+ also have "\<dots> \<le> m * ((n - 1) div m) + (n - 1) mod m" by linarith
+ also have "\<dots> = n - 1" using mod_mult_div_eq by auto
+ finally show ?thesis using \<open>n \<noteq> 0\<close> by auto
+ qed
+ ultimately show ?case by auto
+ qed
+ qed
+ then have "\<Sum>(multiples n m) = \<Sum>{m * q | q. q \<le> (n - 1) div m}"
+ by auto
+ also have "\<dots> = (\<Sum>q \<in> {..(n - 1) div m}. m * q)"
+ by (rule Sum_f) (auto simp add: \<open>m \<noteq> 0\<close> intro: injI)
+ also have "\<dots> = m * \<Sum>{..(n - 1) div m}"
+ using sum_distrib_left[where r = m and f = "\<lambda>k. k"] by auto
+ also have "\<dots> = m * (((n - 1) div m) * ((n - 1) div m + 1) div 2)"
+ proof -
+ {
+ fix n :: nat
+ have "\<Sum>{..n} = n * (n + 1) div 2"
+ by (induction n) auto
+ }
+ then have "\<Sum>{..(n - 1) div m} = ((n - 1) div m) * ((n - 1) div m + 1) div 2" by auto
+ then show ?thesis by auto
+ qed
+ also have "\<dots> = m * ((n - 1) div m) * ((n - 1) div m + 1) div 2"
+ proof -
+ have "m * ((n - 1) div m) * ((n - 1) div m + 1) div 2 = m * (((n - 1) div m) * ((n - 1) div m + 1) div 2) + m * ((n - 1) div m * ((n - 1) div m + 1) mod 2) div 2"
+ using div_mult1_eq[where a = m and b = "((n - 1) div m) * ((n - 1) div m + 1)" and c = 2] by linarith
+ also have "\<dots> = m * (((n - 1) div m) * ((n - 1) div m + 1) div 2)"
+ using Sum_multiples_even by auto
+ finally show ?thesis by auto
+ qed
+ finally show ?thesis .
+qed auto
+
+definition solve001_ap :: "nat \<Rightarrow> nat" where
+ "solve001_ap n = (3 * ((n - 1) div 3) * ((n - 1) div 3 + 1)) div 2 + (5 * ((n - 1) div 5) * ((n - 1) div 5 + 1)) div 2 - (15 * ((n - 1) div 15) * ((n - 1) div 15 + 1)) div 2"
+
+lemma solve001_ap_correctness: "solve001_ap = solve001"
+ unfolding solve001_ap_def using Sum_multiples solve001_multiples by auto
+
+(* Checks if code generation is possible. Commented out, since I probably shouldn't create files
+ on the submission server ^^
+export_code solve001_naive solve001_ap in Haskell
+ module_name ProjectEuler001 file "output/"
+*)
+
+subsection \<open>Project Euler #2: Even Fibonacci numbers\<close>
+
+text \<open>
+ By considering the terms in the Fibonacci sequence whose values do not exceed $N$, find the sum of
+ the even-valued terms.
+\<close>
+
+abbreviation hd2 :: "'a list \<Rightarrow> 'a" where
+ "hd2 \<equiv> hd o tl"
+
+abbreviation last2 :: "'a list \<Rightarrow> 'a" where
+ "last2 \<equiv> last o butlast"
+
+abbreviation butlast2 :: "'a list \<Rightarrow> 'a list" where
+ "butlast2 \<equiv> butlast o butlast"
+
+lemma append_butlast2_last2_last_id:
+ assumes "length xs \<ge> 2"
+ shows "butlast2 xs @ [last2 xs, last xs] = xs"
+proof -
+ have "xs \<noteq> []" using assms by auto
+ then have *: "xs = butlast xs @ [last xs]" by auto
+ also have "\<dots> = butlast2 xs @ [last2 xs] @ [last xs]"
+ proof -
+ have "length (butlast xs) \<ge> 1" using * assms by auto
+ then have "butlast xs \<noteq> []" by fastforce
+ then show ?thesis by auto
+ qed
+ also have "\<dots> = butlast2 xs @ [last2 xs, last xs]" by auto
+ finally show ?thesis by auto
+qed
+
+text \<open>
+ A limited fibonacci sequence.
+\<close>
+
+definition fibonacci_seq :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> bool" where
+ "fibonacci_seq limit n m xs \<longleftrightarrow> n \<noteq> 0 \<and> n \<le> m
+ \<and> hd xs = n \<and> hd2 xs = m
+ \<and> (\<forall>i. i + 2 < length xs \<longrightarrow> xs ! (i + 2) = xs ! i + xs ! (i + 1))
+ \<and> last2 xs + last xs \<ge> limit"
+
+text \<open>
+ Executable version of @{term fibonacci_seq}. Needs termination proof
+\<close>
+
+function fibonaccis :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list" where
+ "fibonaccis limit n m = (if n \<noteq> 0 \<and> n \<le> m \<and> n < limit then n # fibonaccis limit m (n + m) else [])"
+ by auto
+termination
+ by (relation "measures [\<lambda>(limit, n, _). limit - n, \<lambda>(limit, _, m). limit - m]") auto
+
+lemma fibonaccis_Nil_if: "fibonaccis limit n m = [] \<Longrightarrow> n = 0 \<or> n > m \<or> n \<ge> limit"
+ using not_le_imp_less by fastforce
+
+lemma fibonaccis_length:
+ assumes "n \<noteq> 0" "n \<le> m"
+ shows "length (fibonaccis limit m (n + m)) \<le> length (fibonaccis limit n m)"
+ using assms by auto
+
+lemma fibonaccis_lt:
+ assumes "n \<noteq> 0" "n \<le> m"
+ shows "x \<in> set (fibonaccis limit n m) \<Longrightarrow> x < limit"
+proof -
+ assume "x \<in> set (fibonaccis limit n m)"
+ then obtain i where "i < length (fibonaccis limit n m)" "fibonaccis limit n m ! i = x"
+ using in_set_conv_nth[where x = x and xs = "fibonaccis limit n m"] by blast
+ moreover have "i < length (fibonaccis limit n m) \<Longrightarrow> fibonaccis limit n m ! i < limit"
+ using assms
+ proof (induction i arbitrary: n m)
+ case 0
+ then show ?case by auto
+ next
+ case (Suc i)
+ show ?case
+ proof (cases "i = 0")
+ case True
+ then show ?thesis using Suc by auto
+ next
+ case False
+ then have "fibonaccis limit n m ! Suc i = fibonaccis limit m (n+m) ! i"
+ using Suc by auto
+ also have "\<dots> < limit"
+ proof -
+ have "n < limit"
+ proof (rule ccontr, goal_cases False)
+ case False
+ then have "fibonaccis limit n m = []" by auto
+ then show ?case using Suc by auto
+ qed
+ then have "fibonaccis limit n m = n # fibonaccis limit m (n+m)"
+ using Suc by auto
+ then have "i < length (fibonaccis limit m (n+m))" using Suc by auto
+ then show ?thesis using Suc by auto
+ qed
+ finally show ?thesis .
+ qed
+ qed
+ ultimately show "x < limit" by auto
+qed
+
+lemma fibonaccis_fibonacci_seq:
+ assumes "n \<noteq> 0" "n \<le> m" "m < limit"
+ shows "fibonacci_seq limit n m (fibonaccis limit n m)"
+ unfolding fibonacci_seq_def
+ using assms
+proof (safe, goal_cases hd hd2 seq complete)
+ case hd
+ then show ?case by auto
+next
+ case hd2
+ then show ?case by auto
+next
+ case (seq i)
+ then show ?case
+ proof (induction i arbitrary: n m)
+ case 0
+ then show ?case by auto
+ next
+ case (Suc i)
+ then have "fibonaccis limit n m ! (Suc i + 2) = fibonaccis limit m (n + m) ! (i + 2)"
+ by auto
+ also have "\<dots> = fibonaccis limit m (n + m) ! i + fibonaccis limit m (n + m) ! (i + 1)"
+ proof (rule Suc.IH)
+ show "0 < m" using Suc by auto
+ next
+ show "n + m < limit"
+ proof (rule fibonaccis_lt[where n = n and m = m])
+ show "n \<noteq> 0" using Suc by auto
+ next
+ show "n \<le> m" using Suc by auto
+ next
+ have "fibonaccis limit n m ! 2 = n + m"
+ using Suc by auto
+ have "2 < length (fibonaccis limit n m)"
+ proof -
+ have "2 \<le> Suc i + 2" by auto
+ also have "\<dots> < length (fibonaccis limit n m)" using Suc by auto
+ finally show ?thesis .
+ qed
+ then have "fibonaccis limit n m ! 2 \<in> set (fibonaccis limit n m)"
+ using nth_mem[where n = 2 and xs = "fibonaccis limit n m"] by blast
+ moreover have "fibonaccis limit n m ! 2 = n + m" using Suc by auto
+ ultimately show "n + m \<in> set (fibonaccis limit n m)"
+ using back_subst[where a = "fibonaccis limit n m ! 2" and
+ b = "n+m" and
+ P = "\<lambda>x. x \<in> set (fibonaccis limit n m)"]
+ by blast
+ qed
+ next
+ show "i + 2 < length (fibonaccis limit m (n + m))"
+ using Suc fibonaccis_length by auto
+ qed (auto simp add: Suc)
+ also have "\<dots> = fibonaccis limit n m ! (Suc i) + fibonaccis limit n m ! (Suc i + 1)"
+ using Suc by auto
+ finally show ?case .
+ qed
+next
+ case complete
+ let ?xs = "fibonaccis limit n m"
+
+ have "fibonaccis limit (last2 ?xs + last ?xs) (last ?xs + (last2 ?xs + last ?xs)) = [] \<and> last2 ?xs \<noteq> 0"
+ using complete
+ proof (induction ?xs arbitrary: n m rule: length_induct)
+ case 1
+
+
+ let ?xs = "fibonaccis limit n m"
+ let ?ys = "fibonaccis limit m (n + m)"
+
+ {
+ fix n m
+ assume assms: "length (fibonaccis limit n m) < length ?xs" "n \<le> m" "m < limit" "0 < n"
+
+ have aux1: "\<And>A P x. (\<forall>x. A x \<longrightarrow> P x) \<Longrightarrow> A x \<Longrightarrow> P x"
+ by auto
+ have aux2: "\<And>A P n m. (\<forall>n m. A n m \<longrightarrow> P n m) \<Longrightarrow> A n m \<Longrightarrow> P n m"
+ by auto
+ have "fibonaccis limit
+ (last2
+ (fibonaccis limit n m) +
+ last
+ (fibonaccis limit n m))
+ (last
+ (fibonaccis limit n m) +
+ (last2
+ (fibonaccis limit n m) +
+ last
+ (fibonaccis limit n m))) =
+ [] \<and>
+ last2
+ (fibonaccis limit n m) \<noteq>
+ 0" using aux2[OF aux1[OF 1(1) assms(1)], of n m] assms by blast
+ } note IH = this
+
+ have "length ?xs \<ge> 2" using 1 by auto
+
+ have *: "?xs = n # m # fibonaccis limit (n + m) (m + (n + m))" using 1 by auto
+
+ show ?case
+ proof (cases "length ?xs = 2")
+ case True
+ moreover have "?xs = n # m # fibonaccis limit (n + m) (m + (n + m))" using 1 by auto
+ ultimately have "?xs = [n,m]" "fibonaccis limit (n + m) (m + (n + m)) = []" by auto
+ moreover then have "last2 ?xs = n" "last ?xs = m"
+ by auto
+ ultimately show ?thesis using 1 by auto
+ next
+ case False
+ have "fibonaccis limit (last2 ?ys + last ?ys) (last ?ys + (last2 ?ys + last ?ys)) = [] \<and> last2 ?ys \<noteq> 0"
+ proof (rule IH)
+ show "length ?ys < length ?xs" using 1 by auto
+ next
+ show "m \<le> n + m" using 1 by auto
+ next
+ show "n + m < limit"
+ proof (rule ccontr, goal_cases False')
+ case False'
+ then show ?case using False 1 by auto
+ qed
+ next
+ show "0 < m" using 1 by auto
+ qed
+ then have "fibonaccis limit (last2 ?ys + last ?ys) (last ?ys + (last2 ?ys + last ?ys)) = []" "last2 ?ys \<noteq> 0"
+ by blast+
+ moreover have "last2 ?xs = last2 ?ys"
+ proof -
+ have "last2 ?xs = last (butlast (n # fibonaccis limit m (n + m)))"
+ using 1 by auto
+ also have "\<dots> = last (n # butlast (fibonaccis limit m (n + m)))"
+ using 1 by auto
+ also have "\<dots> = last2 ?ys"
+ proof -
+ have "length ?xs \<ge> 3" using \<open>length ?xs \<ge> 2\<close> False by auto
+ moreover have "?xs = n # fibonaccis limit m (n + m)" using 1 by auto
+ ultimately have "length (fibonaccis limit m (n + m)) \<ge> 2" by auto
+ then have "length (butlast (fibonaccis limit m (n + m))) > 0" by auto
+ then show ?thesis by auto
+ qed
+ finally show ?thesis .
+ qed
+ moreover have "last ?xs = last ?ys"
+ using 1 by auto
+ ultimately show ?thesis
+ by auto
+ qed
+ qed
+ then have *: "fibonaccis limit (last2 ?xs + last ?xs) (last ?xs + (last2 ?xs + last ?xs)) = []"
+ and "last2 ?xs \<noteq> 0"
+ by blast+
+ then show ?case using fibonaccis_Nil_if[OF *] by auto
+qed
+
+definition solve002 :: "nat \<Rightarrow> nat" where
+ "solve002 n = sum_list (filter even (fibonaccis n 1 2))"
+
+(*
+export_code solve002 in Haskell
+ module_name ProjectEuler002 file "output/"
+*)
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Keinholz_Jonas_jonas.keinholz@tum.de_658/user_error_log.txt Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,2 @@
+Using temporary directory '/tmp/tmp.TkpaSdA9RU'
+Exercise path '/home/admin/suites/exbeoriginal' not found
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Exercises/hwsubm/OR/Kurz_Friedrich_friedrich.kurz@tum.de_646/be_original.thy Sat Jul 29 19:34:29 2017 +0200
@@ -0,0 +1,1332 @@
+(*\<Rightarrow> Simon*)
+theory be_original
+ imports
+ Main
+ "~~/src/HOL/Matrix_LP/Matrix"
+ "~~/src/HOL/Matrix_LP/SparseMatrix"
+ "../../../Public/Thys/Heap_Prelim"
+begin
+
+text {* \ExerciseSheet{11}{7.~7.~2017} *}
+text \<open>
+ \NumHomework{Be Original!}{28.~7.~2017}
+
+ Develop a nice Isabelle formalization yourself!
+
+ \<^item> This homework is for 3 weeks, and will yield 15 points + 15 bonus points.
+ \<^item> You may develop a formalization from all areas, not only functional data structures.
+ \<^item> Set yourself a time frame and some intermediate/minimal goals.
+ Your formalization needs not be universal and complete after 3 weeks.
+ \<^item> You are welcome to discuss the realizability of your project with the tutor!
+ \<^item> In case you should need inspiration to find a project: Sparse matrices, skew binary numbers,
+ arbitrary precision arithmetic (on lists of bits), interval data structures (e.g. interval lists),
+ spatial data structures (quad-trees, oct-trees), Fibonacci heaps, etc.
+
+\<close>
+
+text \<open>
+In this theory simple dense and sparse matrix implementations using lists are being compared.
+
+ \<^item> dense matrix: implemented as a pair containing the dimensions and a list of dense vectors
+ \<^item> dense vector: implemented as a pair containing the vector size and a list of real values
+ \<^item> sparse matrix: implemented as a pair containing the dimensions and a list of sparse vectors
+ \<^item> sparse vector: implemented as a pair containing the vector size and a list pairs of index and a real value (which is never zero)
+
+As the central result it is shown that dense and sparse matrix addition are interchangeable (lemma "condense_preserves_matrix_addition"): i.e.
+
+ condenseM (addM_LIL M N) = addM_SP (condenseM M) (condenseM N)
+
+The proof recurs on the interchangeability of dense and sparse vector addition which could however not be shown completely.
+
+Additional operations (vector scalar multiplication, vector scalar product, vector multiplication for matrices and matrix multiplication) have been defined but could not be proven within the given time frame.
+\<close>
+
+
+
+(*
+ * I. List of List Matrices (LIL Matrices).
+ *)
+section "List of List Matrices (LIL Matrices)"
+subsection "Definitions"
+subsubsection "Ancillary definitions"
+
+fun repeat :: "'a \<Rightarrow> nat \<Rightarrow> 'a list" where
+ "repeat a 0 = []"
+| "repeat a (Suc m) = (a # repeat a m)"
+
+value "take 0 (repeat 1 10)"
+
+definition list_scalarProd :: "real list \<Rightarrow> real list \<Rightarrow> real" where
+ "list_scalarProd l1 l2 \<equiv> sum_list (map (\<lambda> (a, a'). a * a') (zip l1 l2))"
+
+value "list_scalarProd [1..10] [1..5]"
+
+
+
+subsubsection "LIL Matrix and List Vector."
+paragraph "Basic definitions of LIL Matrix and List Vector."
+
+text\<open>A simple list vector.\<close>
+type_synonym vector_LIL = "(nat * real list)"
+
+text\<open>A simple list of list (LIL) matrix.\<close>
+type_synonym matrix_LIL = "((nat * nat) * vector_LIL list)"
+
+
+
+paragraph "LIL Matrix and List Vector Helper Functions."
+
+abbreviation to_list :: "vector_LIL \<Rightarrow> real list" where
+ "to_list V \<equiv> snd V"
+
+abbreviation size_of :: "vector_LIL \<Rightarrow> nat" where
+ "size_of V \<equiv> fst V"
+
+definition column_size_of :: "matrix_LIL \<Rightarrow> nat" where
+ "column_size_of M = (fst \<circ> fst) M"
+
+abbreviation row_size_of :: "matrix_LIL \<Rightarrow> nat" where
+ "row_size_of M \<equiv> (snd \<circ> fst) M"
+
+abbreviation to_rows :: "matrix_LIL \<Rightarrow> vector_LIL list" where
+ "to_rows M \<equiv> snd M"
+
+definition to_columns :: "matrix_LIL \<Rightarrow> vector_LIL list" where
+"to_columns M = (let rows = to_rows M; n = row_size_of M in
+ map (\<lambda> j. (n, map (\<lambda> i. (to_list (rows ! i)) ! j) [0..<n])) [0..<column_size_of M]
+)"
+
+value "to_columns ((3, 3), [(3, [1, 0, 0]), (3, [1, 0, 0]), (3, [1, 0, 0])])"
+
+definition zeroV_LIL :: "nat \<Rightarrow> vector_LIL" where
+ "zeroV_LIL m = (m, repeat 0 m)"
+
+value "zeroV_LIL 3"
+
+definition unitV_LIL :: "nat \<Rightarrow> nat \<Rightarrow> vector_LIL" where
+ "unitV_LIL m i \<equiv> (if i \<ge> m
+ then undefined
+ else (m, (repeat 0 m)[i := 1])
+ )"
+
+value "unitV_LIL 3 0"
+value "unitV_LIL 3 4"
+
+definition zeroM_LIL :: "nat \<Rightarrow> nat \<Rightarrow> matrix_LIL" where
+ "zeroM_LIL m n = ((m, n), map (\<lambda> _. zeroV_LIL m) [1..n])"
+
+value "zeroM_LIL 3 3"
+
+definition identityM_LIL :: "nat \<Rightarrow> matrix_LIL" where
+ "identityM_LIL m = ((m, m), map (\<lambda> i. unitV_LIL m i) [0..<m])"
+
+value "identityM_LIL 3"
+
+
+
+paragraph "LIL Matrix and Vector invariants."
+
+text\<open>List vector invariant.\<close>
+definition invarV_LIL :: "vector_LIL \<Rightarrow> bool" where
+ "invarV_LIL V \<equiv> (length (to_list V) = size_of V)"
+
+text "LIL matrix invariant: the stored and actual number of rows are equal and for all rows the stored and actual number of columns are equal."
+definition invarM_LIL :: "matrix_LIL \<Rightarrow> bool" where
+ "invarM_LIL M \<equiv> (let m = column_size_of M; n = row_size_of M; rows = to_rows M in
+ n = length rows & (\<forall>v\<in>set rows. m = size_of v & invarV_LIL v)
+ )"
+
+
+
+paragraph "Operations on LIL vectors."
+
+definition insertV_LIL :: "vector_LIL \<Rightarrow> nat \<Rightarrow> real \<Rightarrow> vector_LIL" where
+ "insertV_LIL v i x \<equiv> (let k = size_of v in
+ (if i \<ge> k
+ then undefined
+ else (k, (to_list v)[i := x])
+ )
+ )"
+
+definition retrieveV_LIL :: "vector_LIL \<Rightarrow> nat \<Rightarrow> real option" where
+ "retrieveV_LIL v i \<equiv> (let k = size_of v in
+ (if i \<ge> k
+ then None
+ else Some ((to_list v) ! i)
+ )
+ )"
+
+definition addV_LIL :: "vector_LIL \<Rightarrow> vector_LIL \<Rightarrow> vector_LIL" where
+ "addV_LIL v w \<equiv> (let k1 = size_of v; k2 = size_of w in
+ (if k1 = k2
+ then (k1, map (\<lambda> (a, b). a + b) (zip (to_list v) (to_list w)))
+ else undefined
+ )
+ )"
+
+value "addV_LIL (zeroV_LIL 3) (unitV_LIL 3 0)"
+value "addV_LIL (unitV_LIL 3 2) (addV_LIL (unitV_LIL 3 0) (unitV_LIL 3 1))"
+value "addV_LIL (zeroV_LIL 3) (unitV_LIL 4 0)"
+
+definition scalarMultV_LIL :: "vector_LIL \<Rightarrow> real \<Rightarrow> vector_LIL" where
+ "scalarMultV_LIL v a \<equiv> (let k = size_of v in
+ (k, map (op * a) (to_list v))
+ )"
+
+definition scalarProdV_LIL :: "vector_LIL \<Rightarrow> vector_LIL \<Rightarrow> real" where
+ "scalarProdV_LIL v w \<equiv> (if size_of v = size_of w
+ then list_scalarProd (to_list v) (to_list w)
+ else undefined
+ )"
+
+value "scalarProdV_LIL (4, [1, 1, 1, 1]) (4, [1..4])"
+value "scalarProdV_LIL (5, [1..5]) (4, [1..4])"
+
+
+
+paragraph "Operations on LIL matrices."
+
+text \<open>Insert the given element at the given row and column indexes (in case they don't exceed the matrix' row and column sizes).\<close>
+fun insertM_LIL :: "matrix_LIL \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> real \<Rightarrow> matrix_LIL" where
+ "insertM_LIL M i j x = (let m = column_size_of M; n = row_size_of M; l = to_rows M in
+ if i \<ge> n | j \<ge> m
+ then undefined
+ else ((m, n), l[i := insertV_LIL (l ! i) j x])
+ )"
+
+value "insertM_LIL (zeroM_LIL 3 3) 0 0 1"
+value "insertM_LIL (zeroM_LIL 3 3) 1 2 1"
+value "insertM_LIL (zeroM_LIL 3 3) 2 0 1"
+
+definition retrieveM_LIL :: "matrix_LIL \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> real option" where
+ "retrieveM_LIL M i j \<equiv> (let m = column_size_of M; n = row_size_of M in
+ (if i \<ge> n | j \<ge> m
+ then None
+ else retrieveV_LIL (to_rows M ! i) j
+ )
+ )"
+
+value "retrieveM_LIL (zeroM_LIL 3 3) 2 2"
+value "retrieveM_LIL (zeroM_LIL 100 200) 200 3"
+
+text \<open>LIL Matrix Addition.\<close>
+definition addM_LIL :: "matrix_LIL \<Rightarrow> matrix_LIL \<Rightarrow> matrix_LIL" where
+ "addM_LIL M N \<equiv> (let m1 = column_size_of M; m2 = column_size_of N; n1 = row_size_of M; n2 = row_size_of N in
+ (if m1 = m2 & n1 = n2
+ then ((m1, n1), map (\<lambda> (v1, v2). addV_LIL v1 v2) (zip (to_rows M) (to_rows N)))
+ else undefined
+ )
+ )"
+
+value "addM_LIL (identityM_LIL 3) (zeroM_LIL 3 3)"
+value "addM_LIL (identityM_LIL 3) (identityM_LIL 3)"
+value "addM_LIL (identityM_LIL 3) (zeroM_LIL 4 4)"
+
+definition scalarMultM_LIL :: "matrix_LIL \<Rightarrow> real \<Rightarrow> matrix_LIL" where
+ "scalarMultM_LIL M a \<equiv> (let m = column_size_of M; n = row_size_of M in
+ ((m, n), map (\<lambda> v. scalarMultV_LIL v a) (to_rows M))
+ )"
+
+value "scalarMultM_LIL (identityM_LIL 3) 2"
+value "scalarMultM_LIL (insertM_LIL (zeroM_LIL 3 3) 0 0 1) 2"
+
+definition vectorMultM_LIL :: "matrix_LIL \<Rightarrow> vector_LIL \<Rightarrow> vector_LIL" where
+ "vectorMultM_LIL M v = (let m = column_size_of M; k = size_of v in
+ (if m = k
+ then (k, map (scalarProdV_LIL v) (to_rows M))
+ else undefined
+ )
+ )"
+
+value "insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (zeroM_LIL 3 3) 0 0 1) 0 1 2) 0 2 3) 1 0 4) 1 1 5) 1 2 6) 2 0 7) 2 1 8) 2 2 9"
+value "vectorMultM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (zeroM_LIL 3 3) 0 0 1) 0 1 2) 0 2 3) 1 0 4) 1 1 5) 1 2 6) 2 0 7) 2 1 8) 2 2 9) (3, [1, 1, 1])"
+
+definition matrixMultM_LIL :: "matrix_LIL \<Rightarrow> matrix_LIL \<Rightarrow> matrix_LIL" where
+ "matrixMultM_LIL M N \<equiv> (let m1 = column_size_of M; m2 = column_size_of N; n1 = row_size_of M; n2 = row_size_of N in
+ (if m1 = n2
+ then (let columns = to_columns N in ((n1, m2), map (\<lambda> row. (m2, map (scalarProdV_LIL row) columns)) (to_rows M)))
+ else undefined
+ )
+ )"
+
+value "matrixMultM_LIL (identityM_LIL 3) (zeroM_LIL 3 3)"
+value "matrixMultM_LIL (identityM_LIL 3) (identityM_LIL 3)"
+value "matrixMultM_LIL (identityM_LIL 3) ((1, 3), [(1, [1]), (1, [1]), (1, [1])])"
+value "(insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (zeroM_LIL 3 3) 0 0 1) 0 1 2) 0 2 3) 1 0 4) 1 1 5) 1 2 6) 2 0 7) 2 1 8) 2 2 9)"
+value "matrixMultM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (zeroM_LIL 3 3) 0 0 1) 0 1 2) 0 2 3) 1 0 4) 1 1 5) 1 2 6) 2 0 7) 2 1 8) 2 2 9)"
+value "matrixMultM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (insertM_LIL (zeroM_LIL 3 3) 0 0 1) 0 1 2) 0 2 3) 1 0 4) 1 1 5) 1 2 6) 2 0 7) 2 1 8) 2 2 9) (zeroM_LIL 3 3)"
+
+
+
+subsection \<open>Proofs\<close>
+subsubsection \<open>Proofs w.r.t. LIL matrix implementation\<close>
+paragraph \<open>Some ancillary proofs concerning the properties of the helper functions.\<close>
+
+lemma[simp]: "\<lbrakk>M = ((m, n), l)\<rbrakk> \<Longrightarrow> column_size_of M = m"
+ unfolding column_size_of_def
+ by simp
+lemma[simp]: "\<lbrakk>M = ((m, n), l)\<rbrakk> \<Longrightarrow> row_size_of M = n"
+ by simp
+lemma[simp]: "\<lbrakk>V = (m, l)\<rbrakk> \<Longrightarrow> size_of V = m"
+ by simp
+
+text \<open>Proofs concerning properties of `[1..j]` with respect to `length`.\<close>
+thm upto.simps
+lemma length_upto_nonneg: "\<lbrakk>m \<ge> 0\<rbrakk> \<Longrightarrow> length [1..1 + m] = 1 + length [1..m]"
+ apply (induction m)
+ apply (auto simp add: upto_rec2)
+ done
+
+lemma length_upto_to_card: "length [1..m] = card (set [1..m])"
+proof (induction m)
+ case (nonneg n)
+ then show ?case
+ proof (induction n)
+ case 0
+ then show ?case proof -
+ have "length [1..0] = card (set [1..0])"
+ by simp
+ then have "length [] = card (set [])"
+ by simp
+ then have "0 = card {}"
+ by simp
+ then show ?case
+ by simp
+ qed
+ next case (Suc m)
+ then show ?case proof -
+ assume PREM: "length [1..m] = card (set [1..m])"
+ have "card (set [1..(Suc m)]) = (card (set [1..m] \<union> set [Suc m]))"
+ by simp
+ also have "\<dots> = card (set [1..m]) + card {Suc m}" using Finite_Set.card_Un_disjoint
+ by auto
+ also have "\<dots> = length [1..m] + 1"
+ using PREM
+ by simp
+ also have "\<dots> = length [1..(Suc m)]"
+ using length_upto_nonneg
+ by simp
+ finally show ?case
+ by simp
+ qed
+ qed
+ next case (neg n)
+ then show ?case by simp
+ qed
+
+corollary length_upto: "length [1..m] = m"
+ by (simp add: length_upto_to_card)
+
+lemma length_upto_suc:
+ assumes "length [1..m] = m"
+ shows "length [1..1 + m] = 1 + m"
+proof -
+ have "length [1..Suc m] = card (set [1..Suc m])"
+ using length_upto_to_card
+ by blast
+ also have "\<dots> = Suc m"
+ by simp
+ finally show ?thesis
+ by (simp add: length_upto_to_card)
+qed
+
+
+
+text \<open>Example proofs of the respective invariants for zero vector and zero matrix.\<close>
+
+
+paragraph "Proof of operation properties"
+
+(*
+text "Show that insertion preserves the LIL matrix invariant."
+lemma insertm_lil_invar:
+ assumes INVAR: "invarM_LIL M"
+ assumes "i < row_size_of M"
+ assumes "j < column_size_of M"
+ shows "invarM_LIL (insertM_LIL M i j x)"
+proof -
+ obtain m where "m = (column_size_of (insertM_LIL M i j x))" by blast
+ obtain n where "n = (row_size_of (insertM_LIL M i j x))" by blast
+ obtain l where "l = (to_rows (insertM_LIL M i j x))" by blast
+ have "insertM_LIL ((m, n), l) i j x = ((m, n), l[i := (l ! i)[j := x]])"
+ by (smt B \<open>m = column_size_of (insertM_LIL M i j x)\<close> \<open>n = row_size_of (insertM_LIL M i j x)\<close> column_size_of_def comp_def fst_conv insertM_LIL.simps not_less snd_conv)
+ then have "invarM_LIL (insertM_LIL ((m, n), l) i j x) = invarM_LIL ((m, n), l[i := (l ! i)[j := x]])"
+ by simp
+ then show ?thesis
+ by (smt assms column_size_of_def comp_apply in_set_conv_nth insertM_LIL.simps invarM_LIL_def length_list_update nth_list_update_eq nth_list_update_neq prod.sel(1) prod.sel(2))
+qed
+
+text \<open>Show that for a matrix conforming to the invariant and a pair of indexes within the matrix bounds the sequence of insertion and retrieval at the same indexes yields the inserted value (represented as \texttt{Some x}).\<close>
+lemma retrieve_after_insertm:
+ assumes "invarM_LIL M"
+ assumes "i < row_size_of M"
+ assumes "j < column_size_of M"
+ shows "(Some x) = retrieveM_LIL (insertM_LIL M i j x) i j"
+proof -
+ obtain l where "l = to_rows M"
+ by blast
+ obtain m n l' where "((m, n), l') = insertM_LIL M i j x" and "l' = l[i := (l ! i)[j := x]]"
+ by (smt \<open>l = to_rows M\<close> assms(2) assms(3) insertM_LIL.simps not_less)
+ have "retrieveM_LIL (insertM_LIL M i j x) i j = retrieveM_LIL ((m, n), l') i j"
+ by (simp add: \<open>((m, n), l') = insertM_LIL M i j x\<close>)
+ hence "retrieveM_LIL (insertM_LIL M i j x) i j = Some ((l' ! i) ! j)"
+ by (smt \<open>((m, n), l') = insertM_LIL M i j x\<close> assms(2) assms(3) insertM_LIL.simps not_less prod.sel(2) retrieveM_LIL.simps)
+ hence "retrieveM_LIL (insertM_LIL M i j x) i j = Some ((l[i := (l ! i)[j := x]]) ! i ! j)"
+ using \<open>l' = l[i := (l ! i)[j := x]]\<close>
+ by blast
+ hence "retrieveM_LIL (insertM_LIL M i j x) i j = Some x"
+ using assms
+ by (metis \<open>l = to_rows M\<close> invarM_LIL_def nth_list_update nth_mem)
+ thus ?thesis
+ by simp
+qed
+
+text "Proof the matrix size property of matrix multiplication."
+lemma matrixmult_lil_size:
+ assumes "invarM_LIL M"
+ assumes "invarM_LIL N"
+ assumes "column_size_of M = row_size_of N"
+ shows "row_size_of (matrixMultM_LIL M N) = row_size_of M & column_size_of (matrixMultM_LIL M N) = column_size_of N & invarM_LIL (matrixMultM_LIL M N)"
+proof -
+ show ?thesis sorry
+qed
+
+text "Show that matrix multiplication preserves the LIL matrix invariant."
+lemma matrixmult_lil_invar:
+ assumes "invarM_LIL M"
+ assumes "invarM_LIL N"
+ assumes "column_size_of M = row_size_of N"
+ shows "invarM_LIL (matrixMultM_LIL M N)"
+proof -
+ obtain MN where "MN = matrixMultM_LIL M N"
+ by simp
+ have A: "row_size_of MN = row_size_of M"
+ using matrixmult_lil_size assms `MN = matrixMultM_LIL M N`
+ by simp
+ have B: "column_size_of MN = column_size_of N"
+ using matrixmult_lil_size assms `MN = matrixMultM_LIL M N`
+ by simp
+ have "invarM_LIL MN = ((row_size_of MN = length (to_rows MN)) & (\<forall>p'\<in>set (to_rows MN). column_size_of MN = length p'))"
+ using invarM_LIL_def
+ by meson
+ hence "invarM_LIL MN = ((row_size_of M = length (to_rows MN)) & (\<forall>p'\<in>set (to_rows MN). column_size_of MN = length p'))"
+ using A
+ by simp
+ hence "invarM_LIL MN = ((row_size_of M = length (to_rows MN)) & (\<forall>p'\<in>set (to_rows MN). column_size_of N = length p'))"
+ using B
+ by simp
+ thus ?thesis using assms matrixmult_lil_size by simp
+qed
+*)
+
+
+(*
+ * II. Sparse Matrices.
+ *)
+section "Sparse Matrices"
+subsection "Definitions"
+subsubsection "Sparse Matrix and Sparse Vector"
+
+type_synonym listVector_SP = "(nat * real) list"
+type_synonym vector_SP = "(nat * listVector_SP)"
+type_synonym matrix_SP = "((nat * nat) * vector_SP list)"
+
+definition zeroM_SP :: "nat \<Rightarrow> nat \<Rightarrow> matrix_SP" where
+ "zeroM_SP m n = ((m, n), repeat (m, []) n)"
+
+definition zeroV_SP :: "nat \<Rightarrow> vector_SP" where
+ "zeroV_SP m = (m, Nil)"
+
+value "zeroV_SP 3"
+
+definition identityM_SP :: "nat \<Rightarrow> matrix_SP" where
+ "identityM_SP m = ((m, m), map (\<lambda> i. (m, [(i, 1)])) [0..<m])"
+
+value "identityM_SP 3"
+
+definition condenseV :: "vector_LIL \<Rightarrow> vector_SP" where
+ "condenseV v \<equiv> (let k = size_of v; l = to_list v in
+ (k, filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l))
+ )"
+
+value "condenseV (4, [0, 0, 1, 0])"
+value "map condenseV [(4, [1, 0, 0, 0]), (4, [0, 1, 0, 0]), (4, [0, 0, 1, 0]), (4, [0, 0, 0, 1])]"
+
+definition condenseM :: "matrix_LIL \<Rightarrow> matrix_SP" where
+ "condenseM M \<equiv> (let m = column_size_of M; n = row_size_of M; rows = to_rows M in
+ ((m, n), map condenseV rows)
+ )"
+
+
+
+subsubsection "Operations"
+paragraph "Operations on sparse vectors."
+
+definition indexesLV_SP :: "listVector_SP \<Rightarrow> nat list" where
+ "indexesLV_SP l = (map fst l)"
+
+value "indexesLV_SP []"
+value "indexesLV_SP (map (\<lambda> i. (nat i, i)) [1..3])"
+value "indexesLV_SP [(0, 5), (9, 14), (110, 31)]"
+
+definition to_list_SP :: "vector_SP \<Rightarrow> listVector_SP" where
+ "to_list_SP v = snd v"
+
+definition size_of_SP :: "vector_SP \<Rightarrow> nat" where
+ "size_of_SP v \<equiv> fst v"
+
+fun insertV_SP' :: "listVector_SP \<Rightarrow> nat \<Rightarrow> real \<Rightarrow> listVector_SP" where
+ "insertV_SP' [] i x = [(i, x)]"
+| "insertV_SP' ((i', x') # rest) i x = (if i' = i
+ then (i, x) # rest
+ else (if i' > i
+ then ((i, x) # (i', x') # rest)
+ else (i', x') # (insertV_SP' rest i x)
+ )
+ )"
+definition insertV_SP :: "vector_SP \<Rightarrow> nat \<Rightarrow> real \<Rightarrow> vector_SP" where
+ "insertV_SP v i x \<equiv> (let k = size_of_SP v in
+ (if i \<ge> k then undefined else (k, insertV_SP' (to_list_SP v) i x))
+ )"
+
+value "insertV_SP (1, []) 3 1"
+value "insertV_SP (3, []) 2 1"
+value "insertV_SP (3, [(0, 1), (2, 3)]) 1 2"
+value "insertV_SP (3, [(0, 1), (1, 2)]) 2 3"
+
+definition retrieveV_SP :: "vector_SP \<Rightarrow> nat \<Rightarrow> (nat * real) option" where
+ "retrieveV_SP v i \<equiv> (if i \<ge> (size_of_SP v)
+ then undefined
+ else find (\<lambda> (index, x). index = i) (to_list_SP v))"
+
+value "retrieveV_SP (3, []) 1"
+value "retrieveV_SP (3, [(2, 1)]) 2"
+
+text "Helper function for vector addition."
+fun addV_SP' :: "listVector_SP \<Rightarrow> listVector_SP \<Rightarrow> listVector_SP" where
+ "addV_SP' [] [] = []"
+| "addV_SP' [] ((i, x) # rest) = (i, x) # rest"
+| "addV_SP' ((i, x) # rest) [] = (i, x) # rest"
+| "addV_SP' ((i1, x1) # rest1) ((i2, x2) # rest2) = (if i1 = i2
+ then (if 0 = x1 + x2
+ then addV_SP' rest1 rest2
+ else (i1, x1 + x2) # addV_SP' rest1 rest2
+ )
+ else (if i1 > i2
+ then (i2, x2) # addV_SP' ((i1, x1) # rest1) rest2
+ else (i1, x1) # addV_SP' rest1 ((i2, x2) # rest2)
+))"
+definition addV_SP :: "vector_SP \<Rightarrow> vector_SP \<Rightarrow> vector_SP" where
+ "addV_SP v w = (if size_of_SP v = size_of_SP w
+ then (size_of_SP v, addV_SP' (to_list_SP v) (to_list_SP w))
+ else undefined
+ )"
+
+value "addV_SP (zeroV_SP 3) (zeroV_SP 3)"
+value "addV_SP (4, [(0, 1), (2, 1)]) (4, [(1, 1), (3, 1)])"
+value "addV_SP (4, [(1, 1), (3, 1)]) (4, [(0, 1), (2, 1)])"
+value "addV_SP (4, [(0, 1), (1, 1), (2, 1), (3, 1)]) (4, [(0, 1), (1, 1), (2, 1), (3, 1)])"
+
+definition scalarMultV_SP :: "vector_SP \<Rightarrow> real \<Rightarrow> vector_SP" where
+ "scalarMultV_SP v a \<equiv> (size_of_SP v, map (\<lambda> (i, x). (i, x * a)) (to_list_SP v))"
+
+value "scalarMultV_SP (zeroV_SP 3) 100"
+value "scalarMultV_SP (3, [(0, 1), (1, 1), (2,1)]) 2"
+
+text \<open>
+ Helper function for sparse vector product.
+ The product is only calculated when indices match. Otherwise the entry for the smaller index is dropped because due to ordering there cannot be a matching index in the second list vector as the entry for the lowest index is already higher.
+\<close>
+fun scalarProdV_SP' :: "listVector_SP \<Rightarrow> listVector_SP \<Rightarrow> real list" where
+ "scalarProdV_SP' [] _ = []"
+| "scalarProdV_SP' _ [] = []"
+| "scalarProdV_SP' ((i1, x1) # rest1) ((i2, x2) # rest2) = (if i1 = i2
+ then (x1 * x2) # (scalarProdV_SP' rest1 rest2)
+ else (if i1 > i2
+ then scalarProdV_SP' ((i1, x1) # rest1) rest2
+ else scalarProdV_SP' rest1 ((i2, x2) # rest2))
+)"
+definition scalarProdV_SP :: "vector_SP \<Rightarrow> vector_SP \<Rightarrow> real" where
+ "scalarProdV_SP v w \<equiv> (if size_of_SP v = size_of_SP w
+ then sum_list (scalarProdV_SP' (to_list_SP v) (to_list_SP w))
+ else undefined
+ )"
+
+value "scalarProdV_SP (4, [(0,1), (2, 1)]) (4, [(1,1), (3,1)])"
+value "scalarProdV_SP (4, [(0,1), (1, 2), (2, 3), (3, 4)]) (4, [(0, 1), (1,1), (2, 1), (3,1)])"
+
+
+
+paragraph "Operations on sparse matrices."
+
+definition column_size_of_SP :: "matrix_SP \<Rightarrow> nat" where
+ "column_size_of_SP M \<equiv> (fst \<circ> fst) M"
+
+definition row_size_of_SP :: "matrix_SP \<Rightarrow> nat" where
+ "row_size_of_SP M \<equiv> (snd \<circ> fst) M"
+
+definition to_rows_SP :: "matrix_SP \<Rightarrow> vector_SP list" where
+ "to_rows_SP M \<equiv> snd M"
+ thm is_none_simps
+
+text \<open>Insert the given element at the given row and column indexes (in case they don't exceed the matrix' row and column sizes).\<close>
+definition insertM_SP :: "matrix_SP \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> real \<Rightarrow> matrix_SP" where
+ "insertM_SP M i j x \<equiv> (let m = column_size_of_SP M; n = row_size_of_SP M in
+ (if i < m & j < n
+ then (let rows = to_rows_SP M; row' = (insertV_SP (rows ! i) j x) in
+ ((m, n), rows[i:=row'])
+ )
+ else undefined
+ ))"
+
+value "insertM_SP (zeroM_SP 3 3) 0 0 1"
+value "(insertM_SP (insertM_SP (insertM_SP (zeroM_SP 3 3) 0 0 1) 1 0 1) 2 0 1)"
+
+definition retrieveM_SP :: "matrix_SP \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> real option" where
+ "retrieveM_SP M i j \<equiv> (case find ((op = j) \<circ> fst) (to_list_SP (to_rows_SP M ! i)) of None \<Rightarrow> None | Some (i, x) \<Rightarrow> Some x)"
+
+value "retrieveM_SP (identityM_SP 3) 0 1"
+
+value "retrieveM_SP (identityM_SP 3) 1 2"
+value "retrieveM_SP (identityM_SP 3) 0 0"
+value "retrieveM_SP (identityM_SP 3) 1 0"
+value "retrieveM_SP (identityM_SP 3) 2 0"
+
+value "retrieveM_SP (identityM_SP 3) 0 0"
+value "retrieveM_SP (identityM_SP 3) 1 1"
+value "retrieveM_SP (identityM_SP 3) 2 2"
+
+thm filter.simps
+definition to_columns_SP :: "matrix_SP \<Rightarrow> vector_SP list" where
+ "to_columns_SP M \<equiv> (let n = row_size_of_SP M; rows = to_rows_SP M in
+ map (\<lambda> j. (n, filter (\<lambda> (i, x). x \<noteq> 0) (map (\<lambda> i. case retrieveV_SP (rows ! i) j of None \<Rightarrow> (i, 0) | Some (_, x) \<Rightarrow> (i, x)) [0..<n]))) [0..<(column_size_of_SP M)]
+ )"
+
+value "(identityM_SP 3)"
+value "to_columns_SP (identityM_SP 3)"
+value "to_columns_SP (insertM_SP (insertM_SP (insertM_SP (zeroM_SP 3 3) 0 0 1) 1 0 1) 2 0 1)"
+
+
+text "Sparse Matrix Addition."
+definition addM_SP :: "matrix_SP \<Rightarrow> matrix_SP \<Rightarrow> matrix_SP" where
+ "addM_SP M N \<equiv> (let m = column_size_of_SP M; n = row_size_of_SP M in
+ (if m = column_size_of_SP N \<and> n = row_size_of_SP N
+ then ((m, n), map (\<lambda> (row1, row2). addV_SP row1 row2) (zip (to_rows_SP M) (to_rows_SP N)))
+ else undefined))"
+
+value "addM_SP (zeroM_SP 3 3) (zeroM_SP 3 3)"
+value "addM_SP (identityM_SP 3) (identityM_SP 3)"
+
+definition scalarMultM_SP :: "matrix_SP \<Rightarrow> real \<Rightarrow> matrix_SP" where
+ "scalarMultM_SP M a \<equiv> (let m = column_size_of_SP M; n = row_size_of_SP M in
+ (if 0 = a
+ then zeroM_SP m n
+ else ((m, n), map (\<lambda> row. scalarMultV_SP row a) (to_rows_SP M))
+ )
+ )"
+
+value "scalarMultM_SP (identityM_SP 3) 2"
+value "scalarMultM_SP ((3, 2), [(3, [(0, 1), (1, 1), (2, 1)]), (3, [(0, 1), (1, 1), (2, 1)])]) 2"
+
+definition vectorMultM_SP :: "matrix_SP \<Rightarrow> vector_SP \<Rightarrow> vector_SP" where
+ "vectorMultM_SP M v \<equiv> (let m = column_size_of_SP M; n = row_size_of_SP M; k = size_of_SP v in
+ (if m = k
+ then (n, map (\<lambda> (i, row). (i, scalarProdV_SP row v)) (zip [0..<n] (to_rows_SP M)))
+ else undefined
+ )
+ )"
+
+value "map (\<lambda> (i, c). scalarProdV_SP (3, [(0, 1), (1, 1), (2, 1)]) c) (zip [0..<3] [(3, [(0, 1)]), (3, [(2, 1)]), (3, [(2, 1)])])"
+
+definition matrixMultM_SP :: "matrix_SP \<Rightarrow> matrix_SP \<Rightarrow> matrix_SP" where
+ "matrixMultM_SP M N \<equiv> (let m = column_size_of_SP M; n = row_size_of_SP N in
+ (if m = n
+ then (let rows = to_rows_SP M; columns = to_columns_SP N; k = column_size_of_SP M in
+ (
+ (row_size_of_SP M, column_size_of_SP N),
+ map (\<lambda> row.
+ (k, filter (\<lambda> (i, r). 0 \<noteq> r) (map (\<lambda> (i, column). (i, scalarProdV_SP row column )) (zip [0..<k] columns)))
+ ) rows)
+ )
+ else undefined
+ )
+ )"
+
+value "matrixMultM_SP (zeroM_SP 3 3) (identityM_SP 3)"
+value "insertM_SP (zeroM_SP 3 3) 1 0 1"
+value "identityM_SP 3"
+value "matrixMultM_SP (insertM_SP (zeroM_SP 3 3) 1 0 1) (insertM_SP (zeroM_SP 3 3) 0 1 1)"
+value "insertM_SP (zeroM_SP 3 3) 0 2 1"
+value "insertM_SP (zeroM_SP 3 3) 2 1 1"
+value "matrixMultM_SP (insertM_SP (zeroM_SP 3 3) 0 2 1) (insertM_SP (zeroM_SP 3 3) 2 1 1)"
+value "(insertM_SP (insertM_SP (insertM_SP (zeroM_SP 3 3) 0 0 1) 1 0 1) 2 0 1)"
+value "(insertM_SP (insertM_SP (insertM_SP (zeroM_SP 3 3) 0 1 1) 0 1 1) 0 2 1)"
+value "matrixMultM_SP (insertM_SP (insertM_SP (insertM_SP (zeroM_SP 3 3) 0 0 1) 1 0 1) 2 0 1) (insertM_SP (insertM_SP (insertM_SP (zeroM_SP 3 3) 0 0 1) 0 1 1) 0 2 1)"
+
+
+
+subsubsection "Invariants"
+
+text "Sparse vector invariant: the index list is strictly ascending and no index exceeds or is equal to the stored size."
+definition invarV_SP :: "vector_SP \<Rightarrow> bool" where
+ "invarV_SP v \<equiv> (let indexes = indexesLV_SP (to_list_SP v); size = size_of_SP v in
+ strictly_ascending indexes & (\<forall>i\<in>set indexes. i < size)
+ )"
+
+text "Sparse matrix invariant: the number of rows is equal to the stored row size and all vectors in all rows have a stored size equal to the column size and conform to the sparse vector invariant."
+definition invarM_SP :: "matrix_SP \<Rightarrow> bool" where
+ "invarM_SP M \<equiv> (length (to_rows_SP M) = row_size_of_SP M & (\<forall> v \<in> set (to_rows_SP M). size_of_SP v = column_size_of_SP M & invarV_SP v))"
+
+
+
+subsection "Proofs"
+
+
+lemma "invarV_LIL v \<Longrightarrow> sorted [0..<size_of v]"
+ by simp
+lemma "invarV_LIL v \<Longrightarrow> sorted [0..<size_of v] \<Longrightarrow> sorted (map fst (zip [0..<size_of v] (to_list v)))"
+ apply(induction "to_list v")
+ apply(auto)
+ subgoal sledgehammer
+ by (simp add: invarV_LIL_def)
+ done
+
+lemma[simp]: "invarV_LIL v \<Longrightarrow> sorted [0..<size_of v] \<Longrightarrow> sorted (map fst (zip [0..<size_of v] (to_list v))) \<Longrightarrow> sorted (map fst (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<size_of v] (to_list v))))"
+ using sorted_filter
+ by blast
+
+lemma "invarV_LIL v \<Longrightarrow> distinct [0..<size_of v]"
+ by simp
+lemma "invarV_LIL v \<Longrightarrow> distinct [0..<size_of v] \<Longrightarrow> distinct (map fst (zip [0..<size_of v] (to_list v)))"
+ apply(induction "to_list v")
+ apply(auto simp add: distinct_def)
+ subgoal
+ by (simp add: invarV_LIL_def)
+ done
+lemma[simp]:"invarV_LIL v \<Longrightarrow> distinct [0..<size_of v] \<Longrightarrow> distinct (map fst (zip [0..<size_of v] (to_list v))) \<Longrightarrow> distinct (map fst (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<size_of v] (to_list v))))"
+ using distinct_map_filter
+ by blast
+
+lemma condense_preserves_vector_invar_aux_1:
+ assumes "invarV_LIL v"
+ shows "strictly_ascending (indexesLV_SP (to_list_SP (condenseV v)))"
+proof -
+ obtain k l where "k = size_of v" and "l = to_list v"
+ by blast
+ have "indexesLV_SP (to_list_SP (condenseV v)) = indexesLV_SP (to_list_SP ((k, filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l))))"
+ by (metis \<open>k = size_of v\<close> \<open>l = to_list v\<close> condenseV_def)
+ also have "\<dots> = indexesLV_SP (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l))"
+ using to_list_SP_def
+ by auto
+ also have "\<dots> = (map fst (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l)))"
+ using indexesLV_SP_def
+ by blast
+ hence "strictly_ascending (indexesLV_SP (to_list_SP (condenseV v))) = strictly_ascending (map fst (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l)))"
+ by (simp add: calculation)
+ hence "strictly_ascending (indexesLV_SP (to_list_SP (condenseV v))) = sorted (map fst (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l))) & distinct (map fst (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l)))"
+ by (metis (no_types, lifting) \<open>k = size_of v\<close> \<open>l = to_list v\<close> assms diff_zero distinct_map_filter distinct_upt invarV_LIL_def length_upt map_fst_zip strictly_ascending_def)
+ thus ?thesis
+ using \<open>k = size_of v\<close> \<open>l = to_list v\<close> assms invarV_LIL_def
+ by auto
+qed
+
+lemma condense_preserves_vector_invar_aux_2_1:
+ "(set (map fst (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l))) = fst ` set (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l)))"
+ by simp
+
+lemma condense_preserves_vector_invar_aux_2_2: "(set (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l))) \<subseteq> set (zip [0..<k] l)"
+ by auto
+
+lemma condense_preserves_vector_invar_aux_2_3: "fst ` set (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l)) \<subseteq> fst ` set (zip [0..<k] l)"
+ using condense_preserves_vector_invar_aux_2_2
+ by blast
+
+lemma condense_preserves_vector_invar_aux_2_4:
+ assumes "invarV_LIL v"
+ shows "fst ` set (zip [0..<size_of v] (to_list v)) = set [0..<size_of v]"
+ by (metis assms diff_zero invarV_LIL_def length_upt map_fst_zip set_map)
+
+lemma condense_preserves_vector_invar_aux_2:
+ assumes "invarV_LIL v"
+ shows "(set (map fst (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<size_of v] (to_list v))))) \<subseteq> (set [0..<size_of v])"
+proof -
+ obtain k l where "k = size_of v" and "l = to_list v"
+ by blast
+ have "(set (map fst (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l))) = fst ` set (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l)))"
+ by simp
+ hence "(set (map fst (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l))) \<subseteq> fst ` set (zip [0..<k] l))"
+ using condense_preserves_vector_invar_aux_2_1 condense_preserves_vector_invar_aux_2_2
+ by blast
+ hence "(set (map fst (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l)))) \<subseteq> (set [0..<k])"
+ using condense_preserves_vector_invar_aux_2_4
+ using \<open>k = size_of v\<close> \<open>l = to_list v\<close> assms
+ by blast
+ thus ?thesis
+ using \<open>k = size_of v\<close> \<open>l = to_list v\<close>
+ by blast
+qed
+
+lemma condense_preserves_vector_invar_aux_3:
+ assumes "invarV_LIL v"
+ shows "set (indexesLV_SP (to_list_SP (condenseV v))) \<subseteq> set ([0..<size_of_SP (condenseV v)])"
+proof -
+ obtain k l where "k = size_of v" and "l = to_list v"
+ by blast
+ have A: "set ([0..<size_of_SP (condenseV v)]) = set ([0..<k])"
+ by (metis \<open>k = size_of v\<close> condenseV_def fst_conv size_of_SP_def)
+ have B: "set (indexesLV_SP (to_list_SP (condenseV v))) = set (map fst (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l)))"
+ by (metis \<open>k = size_of v\<close> \<open>l = to_list v\<close> condenseV_def indexesLV_SP_def snd_conv to_list_SP_def)
+ hence "(set (indexesLV_SP (to_list_SP (condenseV v))) \<subseteq> set ([0..<size_of_SP (condenseV v)])) = (set (map fst (filter (\<lambda> (i, x). x \<noteq> 0) (zip [0..<k] l))) \<subseteq> set ([0..<k]))"
+ using A
+ by blast
+ thus ?thesis
+ using \<open>k = size_of v\<close> \<open>l = to_list v\<close> assms condense_preserves_vector_invar_aux_2
+ by blast
+qed
+
+lemma condense_preserves_vector_invar_aux_4:
+ assumes "invarV_LIL v"
+ shows "\<forall>i\<in>set (indexesLV_SP (to_list_SP (condenseV v))). i < size_of_SP (condenseV v)"
+proof -
+ fix i
+ have "i\<in>set (indexesLV_SP (to_list_SP (condenseV v))) \<Longrightarrow> i \<in> set ([0..<size_of_SP (condenseV v)])"
+ using assms condense_preserves_vector_invar_aux_3
+ by blast
+ hence "i\<in>set (indexesLV_SP (to_list_SP (condenseV v))) \<Longrightarrow> i < size_of_SP (condenseV v)"
+ using atLeast_upt
+ by blast
+ thus ?thesis
+ using condense_preserves_vector_invar_aux_3 assms atLeast_upt
+ by blast
+qed
+
+lemma condense_preserves_vector_invar:
+ assumes "invarV_LIL v"
+ shows "invarV_SP (condenseV v)"
+proof -
+ have A: "invarV_LIL v \<Longrightarrow> strictly_ascending (indexesLV_SP (to_list_SP (condenseV v)))"
+ by (simp add: condense_preserves_vector_invar_aux_1)
+ have B: "invarV_LIL v \<Longrightarrow> \<forall>i\<in>set (indexesLV_SP (to_list_SP (condenseV v))). i < size_of_SP (condenseV v)"
+ using condense_preserves_vector_invar_aux_4
+ by blast
+ thus ?thesis using assms A B
+ by (simp add: invarV_SP_def)
+qed
+
+lemma[simp]:
+ assumes "invarM_LIL M"
+ shows "\<forall>v\<in>set(to_rows M). invarV_LIL v"
+proof -
+ fix v
+ have "v\<in>set(to_rows M) \<Longrightarrow> invarV_LIL v" using invarM_LIL_def
+ by (meson assms)
+ thus ?thesis
+ by (meson assms invarM_LIL_def)
+qed
+
+lemma condense_preserves_matrix_invar_aux1:
+ assumes "invarM_LIL M"
+ assumes "v\<in>set(map condenseV (to_rows M))"
+ assumes "\<exists>x\<in>(set (to_rows M)). v = condenseV x"
+ shows "invarV_SP v"
+proof -
+ fix x
+ have "x\<in>(set (to_rows M)) \<Longrightarrow> invarV_LIL x"
+ by (meson assms(1) invarM_LIL_def)
+ hence "x\<in>(set (to_rows M)) \<Longrightarrow> invarV_SP (condenseV x)"
+ using condense_preserves_vector_invar
+ by simp
+ thus ?thesis
+ using assms(1) assms(3) condense_preserves_vector_invar
+ by force
+qed
+
+lemma condense_preserves_matrix_invar_aux2:
+ assumes "invarM_LIL M"
+ shows "\<forall>v\<in>set(map condenseV (to_rows M)). invarV_SP v"
+proof -
+ fix v
+ have "v\<in>set(map condenseV (to_rows M)) \<Longrightarrow> v\<in>condenseV ` (set (to_rows M))"
+ by auto
+ hence "v\<in>set(map condenseV (to_rows M)) \<Longrightarrow> v\<in>{y. \<exists>x\<in>(set (to_rows M)). y = condenseV x}"
+ using Set.image_def
+ by blast
+ hence "v\<in>set(map condenseV (to_rows M)) \<Longrightarrow> (\<exists>x\<in>(set (to_rows M)). v = condenseV x)"
+ using \<open>v \<in> set (map condenseV (to_rows M)) \<Longrightarrow> v \<in> {y. \<exists>x\<in>set (to_rows M). y = condenseV x}\<close>
+ by blast
+ hence "v\<in>set(map condenseV (to_rows M)) \<Longrightarrow> (\<exists>x\<in>(set (to_rows M)). v = condenseV x) \<Longrightarrow> invarV_SP v"
+ using assms condense_preserves_matrix_invar_aux1
+ by blast
+ thus ?thesis
+ using assms condense_preserves_vector_invar
+ by force
+qed
+
+lemma condense_preserves_matrix_invar_aux3:
+ assumes "invarM_LIL M"
+ assumes "v\<in>set(map condenseV (to_rows M))"
+ shows "size_of_SP v = column_size_of M"
+proof -
+ have "v\<in>set(map condenseV (to_rows M)) \<Longrightarrow> v\<in>condenseV ` (set (to_rows M))"
+ by auto
+ hence "v\<in>set(map condenseV (to_rows M)) \<Longrightarrow> v\<in>{y. \<exists>x\<in>(set (to_rows M)). y = condenseV x}"
+ using Set.image_def
+ by blast
+hence "v\<in>set(map condenseV (to_rows M)) \<Longrightarrow> (\<exists>x\<in>(set (to_rows M)). v = condenseV x)"
+ using \<open>v \<in> set (map condenseV (to_rows M)) \<Longrightarrow> v \<in> {y. \<exists>x\<in>set (to_rows M). y = condenseV x}\<close>
+ by blast
+ hence "v\<in>set(map condenseV (to_rows M)) \<Longrightarrow> (\<exists>x\<in>(set (to_rows M)). v = condenseV x) \<Longrightarrow&