theory Demo01
  imports 
    Main
begin

section \<open>Warning\<close>

text \<open>Interactive theorem proving can be addictive!\<close>

section \<open>Beginner's Warning\<close>

text \<open>This theory contains many more elements that have not been
  covered in the first lecture. Just ignore what you do not understand.
  The next line is something you should ignore.\<close>

declare [[names_short]]

text \<open>You can find several manuals and tutorials in the "Documentation"-tab
  on the left. In particular "prog-prove" is an up-to-date introduction to
  Isabelle/HOL. You can also invoke the documentation from the command line, 
  e.g., by invoking:
  
  isabelle doc                 (* list all documentations *)
  isabelle doc prog-prove      (* invoke specific documentation *)
\<close>

section \<open>Documentation and Comments\<close>

text \<open>This is some text that is used to document a theory.
  It is surrounded by cartouches \<open>\<dots>\<close>, which are entered via << and >>.
  Within these texts, one may refer to Isabelle terms, types and theorems,
  e.g., @{typ "bool option"}, @{term "\<forall> x. P x \<longrightarrow> Q x"}, @{thm conjI}.
  You can CTRL-hover/click to get more information on these parts. 

  Document generation will be covered later in this course.\<close>

(* comments are completely ignored and may
   span several lines and even other outer syntax elements, e.g.

lemma "P \<longrightarrow> P" 

   the stated lemma above. *)

section \<open>Examples from the lecture\<close>

subsection \<open>Types\<close>

datatype ('a,'b)tree = Leaf 'a | Node "('a,'b)tree" 'b "('a,'b)tree"
type_synonym ('a)special_tree = "(nat \<times> 'a, 'a list)tree"

datatype 'a list = Nil | Cons 'a "'a list"
type_synonym string = "char list"

text \<open>Here an error occurs: 
  - you see an error in the left- and right margin of this tab
  - navigate the cursor to the problematic line
  - open the "Output"-tab in the bottom to see the error message.\<close>
datatype foo = Bar foo

subsection \<open>Terms\<close>

text \<open>"term" just takes a term as argument and then determines its type.
  the type is visible in the "Output"-tab\<close>

term "map (\<lambda> x :: nat. x + 1) [1, 3]"

text \<open>in the previous term, we use the predefined @{typ "'a List.list"}-type and 
  corresponding @{const map}-function, and these are different from @{typ "'a list"} 
  that is defined above\<close> 


term "(x :: nat) + (y + z) = (x + y) + z \<and> x + y = y + x" 

text \<open>CTRL-click on @{const All} to jump to definition\<close>
term "All (\<lambda> x. P x y)" 

text \<open>Observe the output of the previous line. Usually one directly enters
  universal quantification via \<forall>.\<close>
term "\<forall> x. P x y" 


subsection \<open>Expressiveness of HOL\<close>

text \<open>well-foundedness of a binary relation can be expressed\<close>

type_synonym 'a rel = "'a \<Rightarrow> 'a \<Rightarrow> bool" 
definition "well_founded (R :: 'a rel) = 
  (\<not> (\<exists> f :: nat \<Rightarrow> 'a . \<forall> n :: nat. R (f n) (f (n + 1))))"  

text \<open>the transitive closure of a relation can be expressed\<close>
definition "trans_cl (R :: 'a rel) a b = 
  (\<exists> (f :: nat \<Rightarrow> 'a) (n :: nat). f 0 = a \<and> f n = b \<and> n \<noteq> 0 \<and> 
    (\<forall> i. i < n \<longrightarrow> R (f i) (f (i + 1))))" 

lemma "well_founded (trans_cl R) = well_founded R" 
text \<open>The proof-state is visible in the "Output"-tab, if "Proof state" is enabled\<close>
  oops (* oops aborts a proof attempt *)

text \<open>induction on natural numbers is sound\<close>
lemma "\<forall> P :: nat \<Rightarrow> bool. P 0 \<and> (\<forall> n. P n \<longrightarrow> P (n + 1)) \<longrightarrow> (\<forall> n. P n)" 
  using nat.induct by auto (* do not try to understand proof *)


subsection \<open>Function definitions\<close>

fun append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
  "append Nil ys = ys"
| "append (Cons x xs) ys = Cons x (append xs ys)"

text \<open>The syntax @{typ "'a :: {linorder,numeral} \<Rightarrow> 'a"} is similar 
  to Haskells type-constraints, e.g. @{text "(Ord a, Num a) => a -> a"}.

  Hence, the following definition of partition is restricted to lists of elements that
  provide a linear order so that we can compare elements by @{term "(\<le>)"}.\<close>

fun partition :: "'a :: linorder \<Rightarrow> 'a list \<Rightarrow> 'a list \<times> 'a list" where
  "partition p Nil = (Nil, Nil)" 
| "partition p (Cons x xs) = (let
       part = partition p xs;
       low = fst part;
       high = snd part         
     in if x \<le> p then (Cons x low, high) else (low, Cons x high))" 

text \<open>The definition of qsort is not accepted, since the internal termination
  prover cannot show termination without further help.\<close>
fun qsort :: "'a :: linorder list \<Rightarrow> 'a list" where
  "qsort Nil = Nil"
| "qsort (Cons x xs) = (case partition x xs of
      (low, high) \<Rightarrow> append (qsort low) (Cons x (qsort high)))" 

text \<open>Predicates can be recursive functions, too\<close>
fun is_sorted :: "'a :: linorder list \<Rightarrow> bool" where
  "is_sorted (Cons x (Cons y ys)) = (x \<le> y \<and> is_sorted (Cons y ys))"
| "is_sorted _ = True" 

text \<open>Note that @{term qsort} is still in blue-color, so a free variable. 
  By contrast @{term is_sorted} is a (black) constant.
  
  Here a counter-example to the lemma is immediately detected.
  (there is a hint marked in blue on the left, where details
   are visible in the output-tab.)\<close>
lemma "is_sorted (qsort xs)" 
  oops 

end (* of theory *)