问题
I defined a very simple object-oriented model. The model defines a set of classes and a set of associations.
nonterminal fmaplets and fmaplet
syntax
"_fmaplet" :: "['a, 'a] ⇒ fmaplet" ("_ /↦⇩f/ _")
"_fmaplets" :: "['a, 'a] ⇒ fmaplet" ("_ /[↦⇩f]/ _")
"" :: "fmaplet ⇒ fmaplets" ("_")
"_FMaplets" :: "[fmaplet, fmaplets] ⇒ fmaplets" ("_,/ _")
"_FMapUpd" :: "['a ⇀ 'b, fmaplets] ⇒ 'a ⇀ 'b" ("_/'(_')" [900, 0] 900)
"_FMap" :: "fmaplets ⇒ 'a ⇀ 'b" ("(1[_])")
syntax (ASCII)
"_fmaplet" :: "['a, 'a] ⇒ fmaplet" ("_ /|->f/ _")
"_fmaplets" :: "['a, 'a] ⇒ fmaplet" ("_ /[|->f]/ _")
translations
"_FMapUpd m (_FMaplets xy ms)" ⇌ "_FMapUpd (_FMapUpd m xy) ms"
"_FMapUpd m (_fmaplet x y)" ⇌ "CONST fmupd x y m"
"_FMap ms" ⇌ "_FMapUpd (CONST fmempty) ms"
"_FMap (_FMaplets ms1 ms2)" ↽ "_FMapUpd (_FMap ms1) ms2"
"_FMaplets ms1 (_FMaplets ms2 ms3)" ↽ "_FMaplets (_FMaplets ms1 ms2) ms3"
datatype classes1 =
Object | Person | Employee | Customer | Project | Task | Sprint
abbreviation "associations ≡ [
STR ''ProjectManager'' ↦⇩f [
STR ''projects'' ↦⇩f (Project, 0::nat, 100::nat),
STR ''manager'' ↦⇩f (Employee, 1, 1)],
STR ''ProjectMember'' ↦⇩f [
STR ''member_of'' ↦⇩f (Project, 0, 100),
STR ''members'' ↦⇩f (Employee, 1, 20)],
STR ''ManagerEmployee'' ↦⇩f [
STR ''line_manager'' ↦⇩f (Employee, 0, 1),
STR ''project_manager'' ↦⇩f (Employee, 0, 100),
STR ''employees'' ↦⇩f (Employee, 3, 7)],
STR ''ProjectCustomer'' ↦⇩f [
STR ''projects'' ↦⇩f (Project, 0, 100),
STR ''customer'' ↦⇩f (Customer, 1, 1)],
STR ''ProjectTask'' ↦⇩f [
STR ''project'' ↦⇩f (Project, 1, 1),
STR ''tasks'' ↦⇩f (Task, 0, 100)],
STR ''SprintTaskAssignee'' ↦⇩f [
STR ''sprint'' ↦⇩f (Sprint, 0, 10),
STR ''tasks'' ↦⇩f (Task, 0, 5),
STR ''assignee'' ↦⇩f (Employee, 0, 1)]]"
I defined also a class_roles
predicate which relates a class to a set of association ends navigable from this class:
lemma fmember_code_predI [code_pred_intro]:
"x |∈| xs" if "Predicate_Compile.contains (fset xs) x"
using that by (simp add: Predicate_Compile.contains_def fmember.rep_eq)
code_pred fmember
by (simp add: Predicate_Compile.contains_def fmember.rep_eq)
definition "assoc_end_class ≡ fst"
inductive assoc_refer_class where
"role |∈| fmdom ends ⟹
fmlookup ends role = Some end ⟹
assoc_end_class end = 𝒞 ⟹
assoc_refer_class ends 𝒞 role"
code_pred [show_modes] assoc_refer_class .
inductive class_roles where
"assoc |∈| fmdom assocs ⟹
fmlookup assocs assoc = Some ends ⟹
assoc_refer_class ends 𝒞 from ⟹
role |∈| fmdom ends ⟹
fmlookup ends role = Some end ⟹
role ≠ from ⟹
class_roles assocs 𝒞 assoc from role"
code_pred [show_modes] class_roles .
values "{(x, y, z, a). class_roles associations x y z a}"
This predicate can be evaluated very fast (please see the last line above).
I need to prove that all association ends are unique for each class. For simplicity I'm trying to prove it for the Employee
class:
lemma fmupd_to_rhs:
"fmupd k x xm = y ⟷ y = fmupd k x xm"
by auto
lemma class_roles_unique:
"class_roles associations Employee assoc1 from role ⟹
class_roles associations Employee assoc2 from role ⟹ assoc1 = assoc2"
apply (erule class_roles.cases; erule class_roles.cases;
erule assoc_refer_class.cases; erule assoc_refer_class.cases)
unfolding fmupd_to_rhs
apply (simp)
apply (elim disjE)
apply auto[1]
apply auto[1]
apply auto[1]
(* And so on... Proving of each case is very slow *)
The problem is that it's very slow. Is it possible to simplify a class_roles
predicate using lemmas generated by code_pred
? Or could you suggest a better approach to prove this lemma?
回答1:
The code_pred
command generates equations for class_roles
, one for each inferred mode, and values
uses them. The theorem class_roles.equation
shows them all. If you want to use them to prove your lemma, you must first transform goal or the lemma statement such that one of the generated class_role_...
constants appears. Doing this manually is pretty cumbersome.
You get much better automation if you let the predicate compiler do this transformation for you. Since the lemma contains universally quantified variables (assoc1
, assoc2
, from
, and role
), I recommend that you define the negation of the lemma statement as an inductive predicate, as the negation turns the universal quantifier into an existential, which is modelled by a free variable in the assumptions. Then, you can use the proof method eval
to do the heavy work:
inductive foo where
"foo" if
"class_roles associations Employee assoc1 from role"
"class_roles associations Employee assoc2 from role"
"assoc1 ≠ assoc2"
code_pred foo .
lemma class_roles_unique:
assumes "class_roles associations Employee assoc1 from role"
and "class_roles associations Employee assoc2 from role"
shows "assoc1 = assoc2"
proof -
have "¬ foo" by eval
with assms show ?thesis by(simp add: foo.simps)
qed
Note that eval
uses code generation and evaluation in PolyML, so it computes the result rather than proves it. That is, evaluation is not checked by Isabelle's kernel. The related proof method code_simp
goes through the kernel, but it does not work out of the box in this example because the code setup for String.asciis_of_literals
is missing in Isabelle2018.
The following lemmas provide the missing code equations for literal strings, but code_simp
is very slow with literal strings (normalization
is a bit faster, but not checked by Isabelle's kernel either).
definition dup_bit :: "bool ⇒ integer ⇒ integer" where
"dup_bit b i = i + i + (if b then 1 else 0)"
lemma dup_bit_code [code]:
"dup_bit True 0 = 1"
"dup_bit False 0 = 0"
"dup_bit True (Code_Numeral.Pos n) = Code_Numeral.Pos (num.Bit1 n)"
"dup_bit False (Code_Numeral.Pos n) = Code_Numeral.Pos (num.Bit0 n)"
"dup_bit True (Code_Numeral.Neg n) = - Code_Numeral.sub (num.Bit0 n) Num.One"
"dup_bit False (Code_Numeral.Neg n) = Code_Numeral.Neg (num.Bit0 n)"
by(simp_all add: dup_bit_def Code_Numeral.sub_def nat_of_num_add num_eq_iff)
(metis diff_numeral_special(1) numeral_Bit0 numeral_plus_numeral sub_num_simps(2))
fun integer_of_bits :: "bool list ⇒ integer" where
"integer_of_bits [] = 0"
| "integer_of_bits (b # bs) = dup_bit b (integer_of_bits bs)"
lemma asciis_of_literal_code [code]:
"String.asciis_of_literal (STR '''') = []"
"String.asciis_of_literal (String.Literal b0 b1 b2 b3 b4 b5 b6 s) =
integer_of_bits [b0, b1, b2, b3, b4, b5, b6] # String.asciis_of_literal s"
including literal.lifting by(transfer; simp add: dup_bit_def; fail)+
来源:https://stackoverflow.com/questions/54971177/how-to-simplify-an-inductive-predicate-by-evaluation