How to define a linear ordering on a type? - isabelle

I'm trying to define a conjunction function for 4-valued logic (false, true, null, and error). In my case the conjunction is equivavlent to min function on linear order false < error < null < true.
datatype bool4 = JF | JT | BN | BE
instantiation bool4 :: linear_order
begin
fun leq_bool4 :: "bool4 ⇒ bool4 ⇒ bool" where
"leq_bool4 JF b = True"
| "leq_bool4 BE b = (b = BE ∨ b = BN ∨ b = JT)"
| "leq_bool4 BN b = (b = BN ∨ b = JT)"
| "leq_bool4 JT b = (b = JT)"
instance proof
fix x y z :: bool4
show "x ⊑ x"
by (induct x) simp_all
show "x ⊑ y ⟹ y ⊑ z ⟹ x ⊑ z"
by (induct x; induct y) simp_all
show "x ⊑ y ⟹ y ⊑ x ⟹ x = y"
by (induct x; induct y) simp_all
show "x ⊑ y ∨ y ⊑ x"
by (induct x; induct y) simp_all
qed
end
definition and4 :: "bool4 ⇒ bool4 ⇒ bool4" where
"and4 a b ≡ minimum a b"
I guess there must be an easier way to define a linear order in Isabelle HOL. Could you suggest a simplification of the theory?

You can use the "Datatype_Order_Generator" AFP entry.
Then it's as simple as importing "$AFP/Datatype_Order_Generator/Order_Generator" and declaring derive linorder "bool4". Note that the constructors must be declared in the order you want them when defining your datatype.
Details on how to download and use the AFP locally can be found here.

Related

Isabelle Failed to refine any pending goal during instantiation

datatype 'a list = Cons 'a "'a list" | Nil
instantiation list :: (order) order
begin
fun less_eq_list :: "'a list ⇒ 'a list ⇒ bool" where
"less_eq_list Nil Nil = True" |
"less_eq_list (Cons _ _) Nil = True" |
"less_eq_list Nil (Cons _ _) = False" |
"less_eq_list (Cons _ a) (Cons _ b) = less_eq_list a b"
instance
proof
fix x y:: "'a list"
show "x ≤ x"
apply(induct_tac x)
apply(auto)
done
(* at this point the state is
show x ≤ x
Successful attempt to solve goal by exported rule:
?x2 ≤ ?x2
proof (state)
this:
x ≤ x
goal (3 subgoals):
1. ⋀x y. (x < y) = (x ≤ y ∧ ¬ y ≤ x)
2. ⋀x y z. x ≤ y ⟹ y ≤ z ⟹ x ≤ z
3. ⋀x y. x ≤ y ⟹ y ≤ x ⟹ x = y
*)
show "(x < y) = (x ≤ y ∧ ¬ y ≤ x)"
(* I get an error here
Failed to refine any pending goal
Local statement fails to refine any pending goal
Failed attempt to solve goal by exported rule:
(?x2 < ?y2) = (?x2 ≤ ?y2 ∧ ¬ ?y2 ≤ ?x2)
*)
qed
end
What is wrong with this? The proof of "x ≤ x" worked like a charm. Somehow "(x < y) = (x ≤ y ∧ ¬ y ≤ x)" doesn't match any subgoal.
Class order is a subclass of preorder, which in turn is a subclass of ord. Class ord requires you to define both less_eq (≤) and less (<). In your code, you have correctly defined less_eq_list but forgot to define less_list, and that's why you got an error when trying to prove (x < y) = (x ≤ y ∧ ¬ y ≤ x).

Topological filters in Isabelle

I'm studying topological filters in Filter.thy
theory Filter
imports Set_Interval Lifting_Set
begin
subsection ‹Filters›
text ‹
This definition also allows non-proper filters.
›
locale is_filter =
fixes F :: "('a ⇒ bool) ⇒ bool"
assumes True: "F (λx. True)"
assumes conj: "F (λx. P x) ⟹ F (λx. Q x) ⟹ F (λx. P x ∧ Q x)"
assumes mono: "∀x. P x ⟶ Q x ⟹ F (λx. P x) ⟹ F (λx. Q x)"
typedef 'a filter = "{F :: ('a ⇒ bool) ⇒ bool. is_filter F}"
proof
show "(λx. True) ∈ ?filter" by (auto intro: is_filter.intro)
qed
I don't get this definition. It's quite convoluted so I'll simplify it first
The expression
F (λx. P x) could be simplified to F P (using eta reduction of lambda calculus). The predicate 'a ⇒ bool is really just a set 'a set. Similarly ('a ⇒ bool) ⇒ bool should be 'a set set. Then we could rewrite the axioms as
assumes conj: "P ∈ F ∧ Q ∈ F ⟹ Q ∩ P ∈ F"
assumes mono: "P ⊆ Q ∧ P ∈ F ⟹ Q ∈ F"
Now my question is about the True axiom. It is equivalent to
assumes True: "UNIV ∈ F"
This does not match with the definitions of filters that I ever saw.
The axiom should be instead
assumes True: "{} ∉ F" (* the name True is not very fitting anymore *)
The statement UNIV ∈ F is unnecessary because it follows from axiom mono.
So what's up with this definition that Isabelle provides?
The link provided by Javier Diaz has lots of explanations.
Turns out this is a definition of improper filter. The axiom True is necessary and does not follow from mono. If this axiom was missing then F could be defined as
F P = False
or in set-theory notation, F could be an empty set and mono and conj would then be satisfied vacuously.

How to lift lemmas

Here is a sample datatype with a deterministic relation.
datatype ty1 = A | B | C ty1 | D ty1
inductive rel1 where
"rel1 A (C B)"
| "rel1 (C B) (D A)"
lemma rel1_det:
"rel1 x y ⟹ rel1 x z ⟹ y = z"
by (elim rel1.cases; auto)
I'm trying to lift the lemma for the following type:
datatype 'a ty2 = E 'a | F 'a
abbreviation "rel2 ≡ rel_ty2 rel1"
lemma rel2_det:
"rel2 x y ⟹ rel2 x z ⟹ y = z"
apply (cases x; cases y; auto)
apply (metis rel1_det right_uniqueD right_uniqueI ty2.rel_intros(1) ty2.right_unique_rel)
by (metis rel1_det right_uniqueD right_uniqueI ty2.rel_intros(2) ty2.right_unique_rel)
But the proof is very ugly. I guess it could be simplified using a transfer method, a lifting package or something else. Could you suggest how to use it?
I don't know about using the transfer package to prove this; however, it is easy to prove if you use the predicate right_unique from the library and the rules for it that the datatype package gives you for free:
lemma right_unique_rel1: "right_unique rel1"
by (auto simp: right_unique_def elim: rel1.cases)
lemma right_unique_rel2: "right_unique rel2"
by (intro ty2.right_unique_rel right_unique_rel1)

How to replace existentially quantified variable?

Here is a simple type system:
datatype type =
VoidType
| IntegerType
| RealType
| StringType
datatype val =
VoidVal
| IntegerVal int
| RealVal real
| StringVal string
fun type_of :: "val ⇒ type" where
"type_of (VoidVal) = VoidType"
| "type_of (IntegerVal _) = IntegerType"
| "type_of (RealVal _) = RealType"
| "type_of (StringVal _) = StringType"
with type conformance relation:
inductive less_type :: "type ⇒ type ⇒ bool" (infix "<" 65) where
"IntegerType < RealType"
Integer values can be casted to corresponding real values:
inductive cast :: "val ⇒ val ⇒ bool" where
"cast (IntegerVal x) (RealVal x)"
I'm trying to prove the following lemma. If type of a variable x conforms to RealType, then there exists a value y with type RealType and x can be casted to y.
lemma is_castable_to_real:
"type_of x < RealType ⟹ ∃y. type_of y = RealType ∧ cast x y"
apply (rule exI[of _ "RealVal v"])
I can prove the generic lemma using cases tactics:
lemma is_castable:
"type_of x < τ ⟹ ∃y. type_of y = τ ∧ cast x y"
by (cases x; cases τ; auto simp add: less_type.simps cast.simps)
But I'm trying to understand how to treat existential quantifiers in lemmas. So I'm trying to provide a concrete example RealVal v for y:
type_of x < RealType ⟹ ∃v. type_of (RealVal v) = RealType ∧ cast x (RealVal v)
The problem is that I get the following proposition instead:
type_of x < RealType ⟹ type_of (RealVal v) = RealType ∧ cast x (RealVal v)
What is the kind of variable v? Is it universally quantified variable? How to make it existentially quantified one?
To prove an existential, you can give a concrete example.
In your case, this example can be derived from the assumption of the lemma.
lemma is_castable_to_real:
assumes subtype_of_real: "type_of x < RealType"
shows "∃y. type_of y = RealType ∧ cast x y"
proof -
have "type_of x = IntegerType"
using subtype_of_real less_type.cases by blast
from this obtain i where x_def: "x = IntegerVal i"
by (cases x, auto)
(* prove it for concrete example (RealVal i) *)
have "type_of (RealVal i) = RealType ∧ cast x (RealVal i)"
by (auto simp add: x_def cast.intros)
(* From the concrete example, the existential statement follows: *)
thus "∃y. type_of y = RealType ∧ cast x y" ..
qed
If you just use v before obtaining or defining it somehow, the value will be similar to undefined. It has the correct type, but you do not know anything about it.
If you start the proof without the dash (-) Isabelle will use the default tactic and you would get the subgoal type_of ?y = RealType ∧ cast x ?y. Here ?y is a schematic variable and you can later provide any value that was already available before starting the proof. Maybe this is the kind of variable you get for v, but it is still not clear how you got to the last line in your question.

What happens during function proofs

I am trying to proof a property of the icmp6 checksum function (sum 16bit integers, add carry, invert 16bit integer).
I defined the functions in isabelle. (I know my proofs are terrible)
But for some reason, isabelle can't proof something about the icmp_csum function, it wants to have.
When I replace the oops in the paste with done it produces thousands of lines that just says:
"linarith_split_limit exceeded (current value is 9)"
theory Scratch
imports Main Int List
begin
fun norm_helper :: "nat ⇒ nat" where
"norm_helper x = (let y = divide x 65536 in (y + x - y * 65536))"
lemma "x ≥ 65536 ⟹ norm_helper x < x" by simp
lemma h: "norm_helper x ≤ x" by simp
fun normalize :: "nat ⇒ nat" where
"normalize x = (if x ≥ 65536
then normalize (norm_helper x)
else x)"
inductive norm_to :: "nat ⇒ nat ⇒ bool" where
"(x < 65536) ⟹ norm_to x x"
| "norm_to y z ⟹ y = norm_helper x ⟹ norm_to x z"
lemma ne: "norm_to x y ⟹ y = normalize x"
apply (induct x y rule: norm_to.induct) by simp+
lemma i: "norm_to x y ⟹ x ≥ y"
apply (induct x y rule: norm_to.induct) by simp+
lemma l: "norm_to x y ⟹ y < 65536"
apply (induct x y rule: norm_to.induct) by simp+
lemma en: "y = normalize x ⟹ norm_to x y"
apply (induct x rule: normalize.induct)
proof -
fix x :: nat
assume 1: "(x ≥ 65536 ⟹ y = Scratch.normalize (norm_helper x) ⟹ norm_to (norm_helper x) y)"
assume 2: "y = Scratch.normalize x"
show "norm_to x y"
proof (cases "x ≥ 65536")
show "¬ 65536 ≤ x ⟹ norm_to x y"
using norm_to.intros(1)[of x] 2 by simp
{
assume s: "65536 ≤ x"
have d: "y = normalize (norm_helper x)" using 2 s by simp
show "65536 ≤ x ⟹ norm_to x y"
using 1 d norm_to.intros(2)[of "norm_helper x" y x]
by blast
}
qed
qed
lemma "normalize x ≤ x" using en i by simp
lemma n[simp]: "normalize x < 65536" using en l by blast
fun sum :: "nat list ⇒ nat" where
"sum [] = 0"
| "sum (x#xs) = x + sum xs"
fun csum :: "nat list ⇒ nat" where
"csum xs = normalize (sum xs)"
fun invert :: "nat ⇒ nat" where
"invert x = 65535 - x"
lemma c: "csum xs ≤ 65535" using n[of "sum xs"] by simp
lemma ic: "invert (csum xs) ≥ 0" using c[of xs] by blast
lemma asdf:
assumes "xs = ys"
shows "invert (csum xs) = invert (csum ys)"
using HOL.arg_cong[of "csum xs" "csum ys" invert,
OF HOL.arg_cong[of xs ys csum]] assms(1)
by blast
function icmp_csum :: "nat list ⇒ nat" where
"icmp_csum xs = invert (csum xs)"
apply simp
apply (rule asdf)
apply simp
oops
end
I have no idea why there is tracing output from linarith there, but given that your definition is neither recursive nor performs pattern matching, you can write it as a definition:
definition icmp_csum :: "nat list ⇒ nat" where
"icmp_csum xs = invert (csum xs)"
Another possibility is to change invert to a definition instead of a fun. (In general, if it's neither recursive nor performs pattern matching, definition is preferable because it has much less overhead than fun.)
NB, just import Main, not Main Int List.
Edit: An explanation from Tobias Nipkow on the mailing list:
This is a known issue. In the outdated LNCS 2283 you can find a discussion what to do about it in Section 3.5.3 Simplification and Recursive Functions. The gist: don't use "if", use pattern matching or "case". Or disable if_split.

Resources