(* tested with Isabelle 2025 *)

theory Demo03
  imports Main
begin

(* hide existing constants *)
hide_const Nil Cons conj

datatype Nat = Zero | Succ Nat

term Zero
term Succ

thm Nat.simps(1-3)
thm Nat.induct[of P y]
thm Nat.exhaust[of x P]

datatype List = Nil | Cons Nat List

term Nil
term Cons

thm List.simps(1-3)
thm List.induct[of P xs]
thm List.exhaust[of xs P]


(* Non-emptyness is checked internally, and here the check obviously fails *)
datatype Tree = Node Tree Nat Tree


function conj :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
  "conj True True = True"
| "conj False y = False" 
| "conj x False = False" 
        apply pat_completeness (* first prove pattern completeness *)
       apply auto (* then show that all overlaps are harmless *)
  done

(* finally prove termination, which is trivial for conj *)
termination by lexicographic_order

thm conj.simps

(* fun does the above manual proof steps internally, 
   and gets rid of overlaps by ordering the equations *)

fun conj2 :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
  "conj2 True True = True"
| "conj2 False y = False" 
| "conj2 x False = False" 

thm conj2.simps

fun minimum :: "Nat \<Rightarrow> Nat \<Rightarrow> Nat" where
  "minimum Zero y = y" 
| "minimum x Zero = x" 
| "minimum (Succ x) (Succ y) = Succ (minimum x y)" 

thm minimum.simps

(* fun also resolves underspecification *)

fun minlist :: "List \<Rightarrow> Nat" where
  "minlist (Cons x Nil) = x" 
| "minlist (Cons x xs) = minimum x (minlist xs)" 

thm minlist.simps

end