| author | wenzelm | 
| Sun, 03 Sep 2023 16:44:07 +0200 | |
| changeset 78639 | ca56952b7322 | 
| parent 78107 | 0366e49dab85 | 
| permissions | -rw-r--r-- | 
| 61368 | 1  | 
(* Title: HOL/ex/Transfer_Debug.thy  | 
2  | 
Author: Ondřej Kunčar, TU München  | 
|
3  | 
*)  | 
|
4  | 
theory Transfer_Debug  | 
|
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
61368 
diff
changeset
 | 
5  | 
imports Main "HOL-Library.FSet"  | 
| 61368 | 6  | 
begin  | 
7  | 
||
8  | 
(*  | 
|
9  | 
This file demonstrates some of the typical scenarios  | 
|
10  | 
when transfer or transfer_prover does not produce expected results  | 
|
11  | 
and how the user might handle such cases.  | 
|
12  | 
*)  | 
|
13  | 
||
14  | 
(* As an example, we use finite sets. The following command recreates the environment in which  | 
|
15  | 
the type of finite sets was created and allows us to do transferring on this type. *)  | 
|
16  | 
context  | 
|
17  | 
includes fset.lifting  | 
|
18  | 
begin  | 
|
19  | 
||
20  | 
subsection \<open>1. A missing transfer rule\<close>  | 
|
21  | 
||
| 78107 | 22  | 
(* We will simulate the situation in which there is not any transfer rules for fempty. *)  | 
23  | 
declare bot_fset.transfer[transfer_rule del] fempty_transfer[transfer_rule del]  | 
|
| 61368 | 24  | 
|
| 78107 | 25  | 
(* We want to prove the following theorem about fcard by transfer *)  | 
26  | 
lemma "fcard (A |-| {|x|}) \<le> fcard A"
 | 
|
27  | 
apply transfer  | 
|
| 61368 | 28  | 
(*  | 
| 78107 | 29  | 
Transfer complains that it could not find a transfer rule for |-| with a matching transfer  | 
30  | 
   relation. An experienced user could notice that {||} (in {|x|}, which is special syntax for
 | 
|
31  | 
   finsert x {||}) was transferred to {||} by a a default reflexivity transfer rule (because there
 | 
|
32  | 
   was not any genuine transfer rule for {||}) and fcard was transferred to card using the transfer
 | 
|
33  | 
relation pcr_fset. Therefore transfer is looking for a transfer rule for |-| with a transfer  | 
|
34  | 
relation that mixes (=) and pcr_fset.  | 
|
| 61368 | 35  | 
This situation might be confusing because the real problem (a missing transfer rule) propagates  | 
36  | 
during the transferring algorithm and manifests later in an obfuscated way. Fortunately,  | 
|
37  | 
we could inspect the behavior of transfer in a more interactive way to pin down the real problem.  | 
|
38  | 
*)  | 
|
| 78107 | 39  | 
oops  | 
| 61368 | 40  | 
|
| 78107 | 41  | 
lemma "fcard (A |-| {|x|}) \<le> fcard A"
 | 
42  | 
apply transfer_start  | 
|
43  | 
(* Setups 8 goals for 8 transfer rules that have to be found and the result as the 9th goal, which  | 
|
44  | 
gets synthesized to the final result of transferring when we find the 8 transfer rules. *)  | 
|
45  | 
apply transfer_step  | 
|
46  | 
apply transfer_step  | 
|
| 61368 | 47  | 
(* We can see that the default reflexivity transfer rule was applied and |\<in>|  | 
48  | 
was transferred to |\<in>| \<Longrightarrow> there is no genuine transfer rule for |\<in>|. *)  | 
|
| 78107 | 49  | 
apply transfer_step  | 
50  | 
defer  | 
|
51  | 
apply transfer_step  | 
|
52  | 
apply transfer_step  | 
|
53  | 
apply transfer_step  | 
|
54  | 
apply transfer_step  | 
|
55  | 
apply transfer_end  | 
|
| 61368 | 56  | 
oops  | 
57  | 
||
| 78107 | 58  | 
(* We provide a transfer rule for {||}. *)
 | 
59  | 
lemma [transfer_rule]: "pcr_fset A {} {||}"
 | 
|
60  | 
by (rule bot_fset.transfer)  | 
|
| 61368 | 61  | 
|
| 78107 | 62  | 
lemma "fcard (A |-| {|x|}) \<le> fcard A"
 | 
63  | 
apply transfer_start  | 
|
64  | 
apply transfer_step  | 
|
65  | 
apply transfer_step (* The new transfer rule was selected and |\<in>| was transferred to \<in>. *)  | 
|
66  | 
apply transfer_step+  | 
|
67  | 
apply transfer_end  | 
|
68  | 
by (rule card_Diff1_le)  | 
|
| 61368 | 69  | 
|
70  | 
(* Of course in the real life, we would use transfer instead of transfer_start, transfer_step+ and  | 
|
71  | 
transfer_end. *)  | 
|
| 78107 | 72  | 
lemma "fcard (A |-| {|x|}) \<le> fcard A"
 | 
73  | 
by transfer (rule card_Diff1_le)  | 
|
| 61368 | 74  | 
|
75  | 
subsection \<open>2. Unwanted instantiation of a transfer relation variable\<close>  | 
|
76  | 
||
77  | 
(* We want to prove the following fact. *)  | 
|
78  | 
lemma "finite (UNIV :: 'a::finite fset set)"  | 
|
79  | 
apply transfer  | 
|
80  | 
(* Transfer does not do anything here. *)  | 
|
81  | 
oops  | 
|
82  | 
||
83  | 
(* Let us inspect interactively what happened. *)  | 
|
84  | 
lemma "finite (UNIV :: 'a::finite fset set)"  | 
|
85  | 
apply transfer_start  | 
|
86  | 
apply transfer_step  | 
|
87  | 
(*  | 
|
88  | 
Here we can realize that not an expected transfer rule was chosen.  | 
|
89  | 
We stumbled upon a limitation of Transfer: the tool used the rule Lifting_Set.UNIV_transfer,  | 
|
90  | 
which transfers UNIV to UNIV and assumes that the transfer relation has to be bi-total.  | 
|
91  | 
The problem is that at this point the transfer relation is not known (it is represented by  | 
|
92  | 
a schematic variable ?R) and therefore in order to discharge the assumption "bi_total ?R", ?R is  | 
|
| 67399 | 93  | 
instantiated to (=.) If the relation had been known (we wish pcr_fset (=), which is not bi-total),  | 
94  | 
the assumption bi_total pcr_fset (=) could not have been discharged and the tool would have  | 
|
| 61368 | 95  | 
backtracked and chosen Lifting.right_total_UNIV_transfer, which assumes only right-totalness  | 
96  | 
(and pcr_fset is right-total).  | 
|
97  | 
*)  | 
|
98  | 
back back (* We can force the tool to backtrack and choose the desired transfer rule. *)  | 
|
99  | 
apply transfer_step  | 
|
100  | 
apply transfer_end  | 
|
101  | 
by auto  | 
|
102  | 
||
103  | 
(* Of course, to use "back" in proofs is not a desired style. But we can prioritize  | 
|
104  | 
the rule Lifting.right_total_UNIV_transfer by redeclaring it LOCALLY as a transfer rule.  | 
|
105  | 
*)  | 
|
106  | 
lemma "finite (UNIV :: 'a::finite fset set)"  | 
|
107  | 
proof -  | 
|
108  | 
note right_total_UNIV_transfer[transfer_rule]  | 
|
109  | 
show ?thesis by transfer auto  | 
|
110  | 
qed  | 
|
111  | 
||
112  | 
end  | 
|
113  | 
||
114  | 
(* Let us close the environment of fset transferring and add the rule that we deleted. *)  | 
|
115  | 
lifting_forget fset.lifting (* deletes the extra added transfer rule for |\<in>| *)  | 
|
116  | 
declare fmember_transfer[transfer_rule] (* we want to keep parametricity of |\<in>| *)  | 
|
117  | 
||
118  | 
end  |