1(*
2 * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230)
3 *
4 * SPDX-License-Identifier: BSD-2-Clause
5 *)
6
7theory Corres_Adjust_Preconds
8
9imports
10  "Corres_UL"
11
12begin
13
14text \<open>
15Gadget for adjusting preconditions in a corres rule or similar.
16
17Probably only useful for predicates with two or more related
18preconditions, such as corres.
19
20Used to do some_corres_rule[adj_corres some_intro_rule],
21given e.g. some_intro_rule: @{prop "(s, t) \<in> sr \<Longrightarrow> P s \<Longrightarrow> Q t"}
22Will apply this rule to solve @{prop "Q t"} components in either
23precondition or any sub-conjunct, and will then try to put the
24assumptions @{prop "P s"}, @{prop "(s, t) \<in> sr"} into the right
25places. The premises of the rule can be in any given order.
26
27Concrete example at the bottom.
28\<close>
29
30named_theorems corres_adjust_precond_structures
31
32locale corres_adjust_preconds begin
33
34text \<open>Worker predicates. Broadly speaking, a goal
35of the form "preconds ?A ?B ?C ?D ==> P" expects to
36establish P by instantiating ?A, or failing that ?B,
37etc.
38
39A goal of the form finalise_preconds A exists to
40make sure that schematic conjuncts of A are resolved
41to True.\<close>
42definition
43  preconds :: "bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool"
44where
45  "preconds A B C D = (A \<and> B \<and> C \<and> D)"
46
47definition
48  finalise_preconds :: "bool \<Rightarrow> bool"
49where
50  "finalise_preconds A = True"
51
52text \<open>Use a precond directly to establish goal.\<close>
53lemma consume_preconds:
54  "preconds A True True True \<Longrightarrow> A"
55  "preconds True B True True \<Longrightarrow> B"
56  "preconds True True C True \<Longrightarrow> C"
57  "preconds True True True D \<Longrightarrow> D"
58  by (simp_all add: preconds_def)
59
60lemmas consume_preconds_True = consume_preconds(1)[where A=True]
61
62text \<open>For use as a drule, split a set of schematic preconds
63to give two sets that can be instantiated separately.\<close>
64lemma split_preconds_left:
65  "preconds (A \<and> A') (B \<and> B') (C \<and> C') (D \<and> D') \<Longrightarrow> preconds A B C D"
66  "preconds (A \<and> A') (B \<and> B') (C \<and> C') True \<Longrightarrow> preconds A B C True"
67  "preconds (A \<and> A') (B \<and> B') True True \<Longrightarrow> preconds A B True True"
68  "preconds (A \<and> A') True True True \<Longrightarrow> preconds A True True True"
69  by (simp_all add: preconds_def)
70
71lemma split_preconds_right:
72  "preconds (A \<and> A') (B \<and> B') (C \<and> C') (D \<and> D') \<Longrightarrow> preconds A' B' C' D'"
73  "preconds (A \<and> A') (B \<and> B') (C \<and> C') True \<Longrightarrow> preconds A' B' C' True"
74  "preconds (A \<and> A') (B \<and> B') True True \<Longrightarrow> preconds A' B' True True"
75  "preconds (A \<and> A') True True True \<Longrightarrow> preconds A' True True True"
76  by (simp_all add: preconds_def)
77
78text \<open>For use as an erule. Initiate the precond process,
79creating a finalise goal to be solved later.\<close>
80lemma preconds_goal_initiate:
81  "preconds A B C D \<Longrightarrow> (preconds A B C D \<Longrightarrow> Q)
82    \<Longrightarrow> finalise_preconds (A \<and> B \<and> C \<and> D) \<Longrightarrow> Q"
83  by simp
84
85text \<open>Finalise preconds, trying to replace conjuncts with
86True if they are not yet instantiated.\<close>
87lemma finalise_preconds:
88  "finalise_preconds True"
89  "finalise_preconds A \<Longrightarrow> finalise_preconds B \<Longrightarrow> finalise_preconds (A \<and> B)"
90  "finalise_preconds X"
91  by (simp_all add: finalise_preconds_def)
92
93text \<open>Shape of the whole process for application to regular corres goals.\<close>
94lemma corres_adjust_pre:
95  "corres_underlying R nf nf' rs P Q  f f'
96    \<Longrightarrow> (\<And>s s'. (s, s') \<in> R \<Longrightarrow> preconds (P1 s) (Q1 s') True True \<Longrightarrow> P s)
97    \<Longrightarrow> (\<And>s s'. (s, s') \<in> R \<Longrightarrow> preconds (Q2 s') (P2 s) True True \<Longrightarrow> Q s')
98    \<Longrightarrow> corres_underlying R nf nf' rs (\<lambda>s. P1 s \<and> P2 s) (\<lambda>s'. Q1 s' \<and> Q2 s') f f'"
99  apply (erule stronger_corres_guard_imp)
100   apply (simp add: preconds_def)+
101  done
102
103ML \<open>
104
105structure Corres_Adjust_Preconds = struct
106
107val def_intros = @{thms conjI pred_conj_app[THEN iffD2]
108    bipred_conj_app[THEN fun_cong, THEN iffD2]}
109
110(* apply an intro rule, splitting preconds assumptions to
111   provide unique assumptions for each goal. *)
112fun intro_split ctxt intros i =
113  ((resolve_tac ctxt intros
114        THEN_ALL_NEW (TRY o assume_tac ctxt))
115    THEN_ALL_NEW (fn j => (EVERY (replicate (j - i) (dresolve_tac ctxt @{thms split_preconds_left} j)))
116        THEN dresolve_tac ctxt @{thms split_preconds_right} j)) i
117
118fun handle_preconds ctxt intros =
119  TRY o (eresolve_tac ctxt [@{thm preconds_goal_initiate}]
120    THEN' REPEAT_ALL_NEW (eresolve_tac ctxt @{thms consume_preconds_True}
121        ORELSE' intro_split ctxt (intros @ def_intros)
122        ORELSE' eresolve_tac ctxt @{thms consume_preconds})
123    THEN' REPEAT_ALL_NEW (resolve_tac ctxt @{thms finalise_preconds})
124  )
125
126fun mk_adj_preconds ctxt intros rule = let
127    val xs = [rule] RL (Named_Theorems.get ctxt @{named_theorems corres_adjust_precond_structures})
128    val x = case xs of
129        [] => raise THM ("no unifier with corres_adjust_precond_structures", 1, [rule])
130      | xs => hd xs
131  in x
132    |> ALLGOALS (handle_preconds ctxt intros)
133    |> Seq.hd
134    |> Simplifier.simplify (clear_simpset ctxt addsimps @{thms conj_assoc simp_thms(21-22)})
135  end
136
137val setup =
138      Attrib.setup @{binding "adj_corres"}
139          ((Attrib.thms -- Args.context)
140              >> (fn (intros, ctxt) => Thm.rule_attribute [] (K (mk_adj_preconds ctxt intros))))
141          "use argument theorems to adjust a corres theorem."
142
143end
144
145\<close>
146
147end
148
149declare corres_adjust_preconds.corres_adjust_pre[corres_adjust_precond_structures]
150
151setup Corres_Adjust_Preconds.setup
152
153experiment begin
154
155definition
156  test_sr :: "(nat \<times> nat) set"
157where
158  "test_sr = {(x, y). y = 2 * x}"
159
160lemma test_corres:
161  "corres_underlying test_sr nf nf' dc (\<lambda>x. x < 40) (\<lambda>y. y < 30 \<and> y = 6)
162    (modify (\<lambda>x. x + 2)) (modify (\<lambda>y. 10))"
163  by (simp add: corres_underlying_def simpler_modify_def test_sr_def)
164
165lemma test_adj_precond:
166  "(x, y) \<in> test_sr \<Longrightarrow> x = 3 \<Longrightarrow> y = 6"
167  by (simp add: test_sr_def)
168
169ML \<open>
170Corres_Adjust_Preconds.mk_adj_preconds @{context} [@{thm test_adj_precond}] @{thm test_corres}
171\<close>
172
173lemma foo_adj_corres:
174  "corres_underlying test_sr nf nf' dc (\<lambda>s. s < 40 \<and> s = 3) (\<lambda>s'. s' < 30) (modify (\<lambda>x. x + 2))
175       (modify (\<lambda>y. 10))"
176  by (rule test_corres[adj_corres test_adj_precond])
177
178text \<open>A more general demo of what it does.\<close>
179lemma
180  assumes my_corres: "corres_underlying my_sr nf nf' rvrel P Q a c"
181  assumes my_adj: "\<And>s s'. (s,s') \<in> my_sr \<Longrightarrow> P s \<Longrightarrow> Q s'"
182  shows "corres_underlying my_sr nf nf' rvrel (\<lambda>s. P s \<and> P s) (\<lambda>s'. True) a c"
183  apply (rule my_corres[adj_corres my_adj])
184  done
185
186end
187
188end
189