| author | wenzelm | 
| Sat, 10 Aug 2024 12:12:53 +0200 | |
| changeset 80678 | c5c9b4470d06 | 
| parent 75651 | f4116b7a6679 | 
| child 81946 | ee680c69de38 | 
| permissions | -rw-r--r-- | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Library/Code_Abstract_Nat.thy  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
2  | 
Author: Stefan Berghofer, Florian Haftmann, TU Muenchen  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
3  | 
*)  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
4  | 
|
| 60500 | 5  | 
section \<open>Avoidance of pattern matching on natural numbers\<close>  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
7  | 
theory Code_Abstract_Nat  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
8  | 
imports Main  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
9  | 
begin  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
10  | 
|
| 60500 | 11  | 
text \<open>  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
12  | 
When natural numbers are implemented in another than the  | 
| 69593 | 13  | 
conventional inductive \<^term>\<open>0::nat\<close>/\<^term>\<open>Suc\<close> representation,  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
14  | 
it is necessary to avoid all pattern matching on natural numbers  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
15  | 
altogether. This is accomplished by this theory (up to a certain  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
16  | 
extent).  | 
| 60500 | 17  | 
\<close>  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
18  | 
|
| 60500 | 19  | 
subsection \<open>Case analysis\<close>  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
20  | 
|
| 60500 | 21  | 
text \<open>  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
22  | 
Case analysis on natural numbers is rephrased using a conditional  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
23  | 
expression:  | 
| 60500 | 24  | 
\<close>  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
26  | 
lemma [code, code_unfold]:  | 
| 55415 | 27  | 
"case_nat = (\<lambda>f g n. if n = 0 then f else g (n - 1))"  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
28  | 
by (auto simp add: fun_eq_iff dest!: gr0_implies_Suc)  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
29  | 
|
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
30  | 
|
| 60500 | 31  | 
subsection \<open>Preprocessors\<close>  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
32  | 
|
| 60500 | 33  | 
text \<open>  | 
| 69593 | 34  | 
The term \<^term>\<open>Suc n\<close> is no longer a valid pattern. Therefore,  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
35  | 
all occurrences of this term in a position where a pattern is  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
36  | 
expected (i.e.~on the left-hand side of a code equation) must be  | 
| 
56790
 
f54097170704
prefer plain ASCII / latex over not-so-universal Unicode;
 
wenzelm 
parents: 
55757 
diff
changeset
 | 
37  | 
eliminated. This can be accomplished -- as far as possible -- by  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
38  | 
applying the following transformation rule:  | 
| 60500 | 39  | 
\<close>  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
40  | 
|
| 57426 | 41  | 
lemma Suc_if_eq:  | 
42  | 
assumes "\<And>n. f (Suc n) \<equiv> h n"  | 
|
43  | 
assumes "f 0 \<equiv> g"  | 
|
44  | 
shows "f n \<equiv> if n = 0 then g else h (n - 1)"  | 
|
45  | 
by (rule eq_reflection) (cases n, insert assms, simp_all)  | 
|
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
46  | 
|
| 60500 | 47  | 
text \<open>  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
48  | 
The rule above is built into a preprocessor that is plugged into  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
49  | 
the code generator.  | 
| 60500 | 50  | 
\<close>  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
51  | 
|
| 60500 | 52  | 
setup \<open>  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
53  | 
let  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
54  | 
|
| 57427 | 55  | 
val Suc_if_eq = Thm.incr_indexes 1 @{thm Suc_if_eq};
 | 
56  | 
||
| 55757 | 57  | 
fun remove_suc ctxt thms =  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
58  | 
let  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
59  | 
val vname = singleton (Name.variant_list (map fst  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
60  | 
(fold (Term.add_var_names o Thm.full_prop_of) thms []))) "n";  | 
| 59643 | 61  | 
val cv = Thm.cterm_of ctxt (Var ((vname, 0), HOLogic.natT));  | 
| 59582 | 62  | 
val lhs_of = snd o Thm.dest_comb o fst o Thm.dest_comb o Thm.cprop_of;  | 
63  | 
val rhs_of = snd o Thm.dest_comb o Thm.cprop_of;  | 
|
64  | 
fun find_vars ct = (case Thm.term_of ct of  | 
|
| 69593 | 65  | 
(Const (\<^const_name>\<open>Suc\<close>, _) $ Var _) => [(cv, snd (Thm.dest_comb ct))]  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
66  | 
| _ $ _ =>  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
67  | 
let val (ct1, ct2) = Thm.dest_comb ct  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
68  | 
in  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
69  | 
map (apfst (fn ct => Thm.apply ct ct2)) (find_vars ct1) @  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
70  | 
map (apfst (Thm.apply ct1)) (find_vars ct2)  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
71  | 
end  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
72  | 
| _ => []);  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
73  | 
val eqs = maps  | 
| 57426 | 74  | 
(fn thm => map (pair thm) (find_vars (lhs_of thm))) thms;  | 
75  | 
fun mk_thms (thm, (ct, cv')) =  | 
|
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
76  | 
let  | 
| 57426 | 77  | 
val thm' =  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
78  | 
Thm.implies_elim  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
79  | 
(Conv.fconv_rule (Thm.beta_conversion true)  | 
| 60801 | 80  | 
(Thm.instantiate'  | 
| 59586 | 81  | 
[SOME (Thm.ctyp_of_cterm ct)] [SOME (Thm.lambda cv ct),  | 
| 57426 | 82  | 
SOME (Thm.lambda cv' (rhs_of thm)), NONE, SOME cv']  | 
| 57427 | 83  | 
Suc_if_eq)) (Thm.forall_intr cv' thm)  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
84  | 
in  | 
| 57426 | 85  | 
case map_filter (fn thm'' =>  | 
86  | 
SOME (thm'', singleton  | 
|
87  | 
(Variable.trade (K (fn [thm'''] => [thm''' RS thm']))  | 
|
| 59169 | 88  | 
(Variable.declare_thm thm'' ctxt)) thm'')  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
89  | 
handle THM _ => NONE) thms of  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
90  | 
[] => NONE  | 
| 57426 | 91  | 
| thmps =>  | 
92  | 
let val (thms1, thms2) = split_list thmps  | 
|
93  | 
in SOME (subtract Thm.eq_thm (thm :: thms1) thms @ thms2) end  | 
|
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
94  | 
end  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
95  | 
in get_first mk_thms eqs end;  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
96  | 
|
| 59169 | 97  | 
fun eqn_suc_base_preproc ctxt thms =  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
98  | 
let  | 
| 59582 | 99  | 
val dest = fst o Logic.dest_equals o Thm.prop_of;  | 
| 69593 | 100  | 
val contains_suc = exists_Const (fn (c, _) => c = \<^const_name>\<open>Suc\<close>);  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
101  | 
in  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
102  | 
if forall (can dest) thms andalso exists (contains_suc o dest) thms  | 
| 59169 | 103  | 
then thms |> perhaps_loop (remove_suc ctxt) |> (Option.map o map) Drule.zero_var_indexes  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
104  | 
else NONE  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
105  | 
end;  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
106  | 
|
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
107  | 
val eqn_suc_preproc = Code_Preproc.simple_functrans eqn_suc_base_preproc;  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
108  | 
|
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
109  | 
in  | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
110  | 
|
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
111  | 
  Code_Preproc.add_functrans ("eqn_Suc", eqn_suc_preproc)
 | 
| 
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
112  | 
|
| 
69216
 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 
wenzelm 
parents: 
60801 
diff
changeset
 | 
113  | 
end  | 
| 60500 | 114  | 
\<close>  | 
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
115  | 
|
| 74592 | 116  | 
|
| 
75651
 
f4116b7a6679
Move code lemmas for symbolic computation of bit operations on int to distribution.
 
haftmann 
parents: 
74592 
diff
changeset
 | 
117  | 
subsection \<open>Candidates which need special treatment\<close>  | 
| 
 
f4116b7a6679
Move code lemmas for symbolic computation of bit operations on int to distribution.
 
haftmann 
parents: 
74592 
diff
changeset
 | 
118  | 
|
| 
 
f4116b7a6679
Move code lemmas for symbolic computation of bit operations on int to distribution.
 
haftmann 
parents: 
74592 
diff
changeset
 | 
119  | 
lemma drop_bit_int_code [code]:  | 
| 
 
f4116b7a6679
Move code lemmas for symbolic computation of bit operations on int to distribution.
 
haftmann 
parents: 
74592 
diff
changeset
 | 
120  | 
\<open>drop_bit n k = k div 2 ^ n\<close> for k :: int  | 
| 
 
f4116b7a6679
Move code lemmas for symbolic computation of bit operations on int to distribution.
 
haftmann 
parents: 
74592 
diff
changeset
 | 
121  | 
by (fact drop_bit_eq_div)  | 
| 74592 | 122  | 
|
123  | 
lemma take_bit_num_code [code]:  | 
|
124  | 
\<open>take_bit_num n Num.One =  | 
|
125  | 
(case n of 0 \<Rightarrow> None | Suc n \<Rightarrow> Some Num.One)\<close>  | 
|
126  | 
\<open>take_bit_num n (Num.Bit0 m) =  | 
|
127  | 
(case n of 0 \<Rightarrow> None | Suc n \<Rightarrow> (case take_bit_num n m of None \<Rightarrow> None | Some q \<Rightarrow> Some (Num.Bit0 q)))\<close>  | 
|
128  | 
\<open>take_bit_num n (Num.Bit1 m) =  | 
|
129  | 
(case n of 0 \<Rightarrow> None | Suc n \<Rightarrow> Some (case take_bit_num n m of None \<Rightarrow> Num.One | Some q \<Rightarrow> Num.Bit1 q))\<close>  | 
|
130  | 
apply (cases n; simp)+  | 
|
131  | 
done  | 
|
132  | 
||
| 
51113
 
222fb6cb2c3e
factored out shared preprocessor setup into theory Code_Abstract_Nat, tuning descriptions
 
haftmann 
parents:  
diff
changeset
 | 
133  | 
end  |