Theory Predicate_Compile_Tests
theory Predicate_Compile_Tests
imports "HOL-Library.Predicate_Compile_Alternative_Defs"
begin
declare [[values_timeout = 480.0]]
subsection ‹Basic predicates›
inductive False' :: "bool"
code_pred (expected_modes: bool) False' .
code_pred [dseq] False' .
code_pred [random_dseq] False' .
values [expected "{}" pred] "{x. False'}"
values [expected "{}" dseq 1] "{x. False'}"
values [expected "{}" random_dseq 1, 1, 1] "{x. False'}"
value "False'"
inductive True' :: "bool"
where
"True ==> True'"
code_pred True' .
code_pred [dseq] True' .
code_pred [random_dseq] True' .
thm True'.equation
thm True'.dseq_equation
thm True'.random_dseq_equation
values [expected "{()}"] "{x. True'}"
values [expected "{}" dseq 0] "{x. True'}"
values [expected "{()}" dseq 1] "{x. True'}"
values [expected "{()}" dseq 2] "{x. True'}"
values [expected "{}" random_dseq 1, 1, 0] "{x. True'}"
values [expected "{}" random_dseq 1, 1, 1] "{x. True'}"
values [expected "{()}" random_dseq 1, 1, 2] "{x. True'}"
values [expected "{()}" random_dseq 1, 1, 3] "{x. True'}"
inductive EmptyPred :: "'a ⇒ bool"
code_pred (expected_modes: o => bool, i => bool) EmptyPred .
definition EmptyPred' :: "'a ⇒ bool"
where "EmptyPred' = (λ x. False)"
code_pred (expected_modes: o => bool, i => bool) [inductify] EmptyPred' .
inductive EmptyRel :: "'a ⇒ 'b ⇒ bool"
code_pred (expected_modes: o => o => bool, i => o => bool, o => i => bool, i => i => bool) EmptyRel .
inductive EmptyClosure :: "('a ⇒ 'a ⇒ bool) ⇒ 'a ⇒ 'a ⇒ bool"
for r :: "'a ⇒ 'a ⇒ bool"
code_pred
(expected_modes: (o => o => bool) => o => o => bool, (o => o => bool) => i => o => bool,
(o => o => bool) => o => i => bool, (o => o => bool) => i => i => bool,
(i => o => bool) => o => o => bool, (i => o => bool) => i => o => bool,
(i => o => bool) => o => i => bool, (i => o => bool) => i => i => bool,
(o => i => bool) => o => o => bool, (o => i => bool) => i => o => bool,
(o => i => bool) => o => i => bool, (o => i => bool) => i => i => bool,
(i => i => bool) => o => o => bool, (i => i => bool) => i => o => bool,
(i => i => bool) => o => i => bool, (i => i => bool) => i => i => bool)
EmptyClosure .
thm EmptyClosure.equation
consts a' :: 'a
inductive Fact :: "'a ⇒ 'a ⇒ bool"
where
"Fact a' a'"
code_pred (expected_modes: o => o => bool, i => o => bool, o => i => bool, i => i => bool) Fact .
text ‹
The following example was derived from an predicate in the CakeML formalisation provided by Lars Hupel.
›
inductive predicate_where_argument_is_condition :: "bool ⇒ bool"
where
"ck ⟹ predicate_where_argument_is_condition ck"
code_pred predicate_where_argument_is_condition .
text ‹Other similar examples of this kind:›
inductive predicate_where_argument_is_in_equation :: "bool ⇒ bool"
where
"ck = True ⟹ predicate_where_argument_is_in_equation ck"
code_pred predicate_where_argument_is_in_equation .
inductive predicate_where_argument_is_condition_and_value :: "bool ⇒ bool"
where
"predicate_where_argument_is_condition_and_value ck ⟹ ck
⟹ predicate_where_argument_is_condition_and_value ck"
code_pred predicate_where_argument_is_condition_and_value .
inductive predicate_where_argument_is_neg_condition :: "bool ⇒ bool"
where
"¬ ck ⟹ predicate_where_argument_is_neg_condition ck"
code_pred predicate_where_argument_is_neg_condition .
inductive predicate_where_argument_is_neg_condition_and_in_equation :: "bool ⇒ bool"
where
"¬ ck ⟹ ck = False ⟹ predicate_where_argument_is_neg_condition_and_in_equation ck"
code_pred predicate_where_argument_is_neg_condition_and_in_equation .
text ‹Another related example that required slight adjustment of the proof procedure›
inductive if_as_predicate :: "bool ⇒ 'a ⇒ 'a ⇒ 'a ⇒ bool"
where
"condition ⟹ if_as_predicate condition then_value else_value then_value"
| "¬ condition ⟹ if_as_predicate condition then_value else_value else_value"
code_pred [show_proof_trace] if_as_predicate .
inductive zerozero :: "nat * nat => bool"
where
"zerozero (0, 0)"
code_pred (expected_modes: i => bool, i * o => bool, o * i => bool, o => bool) zerozero .
code_pred [dseq] zerozero .
code_pred [random_dseq] zerozero .
thm zerozero.equation
thm zerozero.dseq_equation
thm zerozero.random_dseq_equation
text ‹We expect the user to expand the tuples in the values command.
The following values command is not supported.›
text ‹Instead, the user must type›
values "{(x, y). zerozero (x, y)}"
values [expected "{}" dseq 0] "{(x, y). zerozero (x, y)}"
values [expected "{(0::nat, 0::nat)}" dseq 1] "{(x, y). zerozero (x, y)}"
values [expected "{(0::nat, 0::nat)}" dseq 2] "{(x, y). zerozero (x, y)}"
values [expected "{}" random_dseq 1, 1, 2] "{(x, y). zerozero (x, y)}"
values [expected "{(0::nat, 0:: nat)}" random_dseq 1, 1, 3] "{(x, y). zerozero (x, y)}"
inductive nested_tuples :: "((int * int) * int * int) => bool"
where
"nested_tuples ((0, 1), 2, 3)"
code_pred nested_tuples .
inductive JamesBond :: "nat => int => natural => bool"
where
"JamesBond 0 0 7"
code_pred JamesBond .
values [expected "{(0::nat, 0::int , 7::natural)}"] "{(a, b, c). JamesBond a b c}"
values [expected "{(0::nat, 7::natural, 0:: int)}"] "{(a, c, b). JamesBond a b c}"
values [expected "{(0::int, 0::nat, 7::natural)}"] "{(b, a, c). JamesBond a b c}"
values [expected "{(0::int, 7::natural, 0::nat)}"] "{(b, c, a). JamesBond a b c}"
values [expected "{(7::natural, 0::nat, 0::int)}"] "{(c, a, b). JamesBond a b c}"
values [expected "{(7::natural, 0::int, 0::nat)}"] "{(c, b, a). JamesBond a b c}"
values [expected "{(7::natural, 0::int)}"] "{(a, b). JamesBond 0 b a}"
values [expected "{(7::natural, 0::nat)}"] "{(c, a). JamesBond a 0 c}"
values [expected "{(0::nat, 7::natural)}"] "{(a, c). JamesBond a 0 c}"
subsection ‹Alternative Rules›