Skip to content

Commit b9843ac

Browse files
authored
Merge pull request #1064 from CakeML/remove-define
Remove more val foo = Define
2 parents 5dbca0e + 98c4d6a commit b9843ac

11 files changed

+48
-35
lines changed

basis/TextIOProgScript.sml

+5-3
Original file line numberDiff line numberDiff line change
@@ -16,9 +16,11 @@ val _ = ml_prog_update (open_module "TextIO");
1616

1717
val _ = ml_prog_update open_local_block;
1818

19-
20-
(* val get_buffered_in_def = Define `get_out (InstreamBuffered)` *)
21-
19+
(*
20+
Definition get_buffered_in_def:
21+
get_out (InstreamBuffered)
22+
End
23+
*)
2224

2325
Datatype:
2426
instream = Instream mlstring

basis/fsFFIScript.sml

+5-2
Original file line numberDiff line numberDiff line change
@@ -286,7 +286,8 @@ Definition ffi_close_def:
286286
End
287287

288288
(* given a file descriptor and an offset, returns 0 and update fs or returns 1
289-
* if an error is met val ffi_seek = Define`
289+
* if an error is met
290+
Definition ffi_seek_def:
290291
ffi_seek bytes fs =
291292
do
292293
assert(LENGTH bytes = 2);
@@ -295,7 +296,9 @@ End
295296
return(LUPDATE 0w 0 bytes, fs')
296297
od ++
297298
return (LUPDATE 1w 0 bytes, fs)
298-
od`; *)
299+
od
300+
End
301+
*)
299302
(* -- *)
300303

301304
(* Packaging up the model as an ffi_part *)

characteristic/cfLetAutoScript.sml

+4-2
Original file line numberDiff line numberDiff line change
@@ -91,8 +91,10 @@ Definition HPROP_INJ_def:
9191
End
9292

9393
(* TODO: use that *)
94-
(* val HPROP_FRAME_IMP_def = Define `HPROP_FRAME_IMP H1 H2 Frame <=>
95-
?s. VALID_HEAP s /\ H1 s /\ (H2 * Frame) s`;
94+
(*
95+
Definition HPROP_FRAME_IMP_def:
96+
HPROP_FRAME_IMP H1 H2 Frame <=> ?s. VALID_HEAP s /\ H1 s /\ (H2 * Frame) s
97+
End
9698
9799
Theorem HPROP_FRAME_HCOND:
98100
HPROP_FRAME_IMP H1 (&PF * H2) Frame <=> PF /\ HPROP_FRAME_IMP H1 H2 Frame

compiler/backend/ag32/proofs/ag32_ffi_codeProofScript.sml

+4-4
Original file line numberDiff line numberDiff line change
@@ -2923,10 +2923,10 @@ Proof
29232923
qexists_tac `0` >> simp[]
29242924
QED
29252925

2926-
val ag32_ffi_get_arg_length_loop1_code_def = Define‘
2926+
Definition ag32_ffi_get_arg_length_loop1_code_def:
29272927
ag32_ffi_get_arg_length_loop1_code =
29282928
GENLIST (λi. EL (i + 2) ag32_ffi_get_arg_length_loop_code) 4
2929-
’;
2929+
End
29302930

29312931
val instn = instn0
29322932
(CONV_RULE (RAND_CONV EVAL)
@@ -2976,13 +2976,13 @@ val loop_code_def' = Q.prove(
29762976
val instn = instn0 loop_code_def'
29772977
val combined = combined0 instn gmw
29782978

2979-
val has_n_args_def = Define‘
2979+
Definition has_n_args_def:
29802980
(has_n_args mem a 0 ⇔ T) ∧
29812981
(has_n_args (mem : word32 -> word8) a (SUC n) ⇔
29822982
∃off. mem (a + n2w off) = 0w ∧
29832983
(∀i. i < off ⇒ mem (a + n2w i) ≠ 0w) ∧
29842984
has_n_args mem (a + n2w off + 1w) n)
2985-
’;
2985+
End
29862986

29872987
Theorem ag32_ffi_get_arg_length_loop_code_thm:
29882988
has_n_args s.MEM (s.R 5w) argc ∧ w2n (s.R 6w) ≤ argc ∧

compiler/backend/word_allocScript.sml

+3-2
Original file line numberDiff line numberDiff line change
@@ -835,7 +835,7 @@ Proof
835835
QED
836836

837837
(* Old representation *)
838-
(* val get_clash_sets_def = Define`
838+
(* Definition get_clash_sets_def:
839839
(get_clash_sets (Seq s1 s2) live =
840840
let (hd,ls) = get_clash_sets s2 live in
841841
let (hd',ls') = get_clash_sets s1 hd in
@@ -867,7 +867,8 @@ QED
867867
(*Catchall for cases where we dont have in sub programs live sets*)
868868
(get_clash_sets prog live =
869869
let i_set = union (get_writes prog) live in
870-
(get_live prog live,[i_set]))`
870+
(get_live prog live,[i_set]))
871+
End
871872
*)
872873

873874
(* Potentially more efficient liveset representation for checking / allocation*)

compiler/inference/proofs/envRelScript.sml

+5-3
Original file line numberDiff line numberDiff line change
@@ -26,11 +26,13 @@ Definition convert_env_def:
2626
convert_env s env = MAP (\(x,t). (x, convert_t (t_walkstar s t))) env
2727
End
2828

29-
(* val convert_decls_def = Define `
30-
convert_decls idecls =
29+
(*
30+
Definition convert_decls_def:
31+
convert_decls idecls =
3132
<| defined_mods := set idecls.inf_defined_mods;
3233
defined_types := set idecls.inf_defined_types;
33-
defined_exns := set idecls.inf_defined_exns|>`;
34+
defined_exns := set idecls.inf_defined_exns|>
35+
End
3436
3537
Theorem convert_append_decls:
3638
!decls1 decls2. convert_decls (append_decls decls1 decls2) = union_decls (convert_decls decls1) (convert_decls decls2)

compiler/repl/repl_init_typesScript.sml

+5-2
Original file line numberDiff line numberDiff line change
@@ -33,8 +33,11 @@ val print_types = let
3333
val _ = print "\n"
3434
in () end
3535
val result_tm = eval_res_thm |> concl |> rand |> rand
36-
val def = Define ‘repl_prog_types = ^result_tm’
37-
val result = eval_res_thm |> CONV_RULE (PATH_CONV "rr" (REWR_CONV (GSYM def)))
36+
Definition repl_prog_types_def:
37+
repl_prog_types = ^result_tm
38+
End
39+
val result = eval_res_thm
40+
|> CONV_RULE (PATH_CONV "rr" (REWR_CONV (GSYM repl_prog_types_def)));
3841

3942
Theorem repl_prog_types_thm = result;
4043

examples/queueProgScript.sml

+6-4
Original file line numberDiff line numberDiff line change
@@ -47,14 +47,15 @@ End
4747

4848
val EmptyQueue_exn_def = EVAL ``EmptyQueue_exn v``;
4949

50-
val lqueue_def = Define‘
50+
Definition lqueue_def:
5151
lqueue qels f r els ⇔
5252
f < LENGTH qels ∧ r < LENGTH qels ∧
5353
(f ≤ r ∧
5454
(∃pj rj. qels = pj ++ els ++ rj ∧ LENGTH pj = f ∧
5555
r + LENGTH rj = LENGTH qels) ∨
5656
r ≤ f ∧ (∃p s mj. qels = s ++ mj ++ p ∧ els = p ++ s ∧
57-
r = LENGTH s ∧ f = r + LENGTH mj))’;
57+
r = LENGTH s ∧ f = r + LENGTH mj))
58+
End
5859

5960
Theorem lqueue_empty:
6061
i < LENGTH xs ⇒ lqueue xs i i []
@@ -114,14 +115,15 @@ QED
114115
operations can be expressed in terms of the abstract value
115116
*)
116117

117-
val QUEUE_def = Define‘
118+
Definition QUEUE_def:
118119
QUEUE A sz els qv ⇔
119120
SEP_EXISTS av fv rv cv qelvs.
120121
REF qv (Conv NONE [av;fv;rv;cv]) *
121122
ARRAY av qelvs *
122123
& (0 < sz ∧ NUM (LENGTH els) cv ∧
123124
∃qels f r. LIST_REL A qels qelvs ∧ NUM f fv ∧ NUM r rv ∧
124-
lqueue qels f r els ∧ LENGTH qels = sz)’;
125+
lqueue qels f r els ∧ LENGTH qels = sz)
126+
End
125127
(*
126128
type_of “QUEUE”;
127129
*)

semantics/cmlPtreeConversionScript.sml

+3-7
Original file line numberDiff line numberDiff line change
@@ -831,8 +831,7 @@ Definition letFromPat_def:
831831
| _ => Mat rhs [(p,body)]
832832
End
833833

834-
local
835-
val ptree_Expr_quotation = `
834+
Definition ptree_Expr_def:
836835
ptree_Expr ent (Lf _) = NONE
837836
ptree_Expr ent (Nd (nt,loc) subs) =
838837
do
@@ -1217,11 +1216,8 @@ local
12171216
seq <- ptree_Eseq seq_pt;
12181217
SOME(e::seq)
12191218
od
1220-
| _ => NONE)`
1221-
in
1222-
1223-
val ptree_Expr_def = Define ptree_Expr_quotation
1224-
end
1219+
| _ => NONE)
1220+
End
12251221

12261222
Definition ptree_StructName_def:
12271223
ptree_StructName (Lf _) = NONE

semantics/proofs/cmlPtreeConversionPropsScript.sml

+4-4
Original file line numberDiff line numberDiff line change
@@ -16,21 +16,21 @@ val _ = option_monadsyntax.temp_add_option_monadsyntax()
1616

1717
(* first, capture those types that we expect to be in the range of the
1818
conversion *)
19-
val user_expressible_tyname_def = Define‘
19+
Definition user_expressible_tyname_def:
2020
(user_expressible_tyname (Short s) ⇔ T) ∧
2121
(user_expressible_tyname (Long m (Short s)) ⇔ T) ∧
2222
(user_expressible_tyname _ ⇔ F)
23-
’;
23+
End
2424
val _ = augment_srw_ss [rewrites [user_expressible_tyname_def]]
2525

2626
Overload ND[local] = “λn. Nd (mkNT n, ARB)”
2727
Overload LF[local] = “λt. Lf (TOK t, ARB)”
2828

29-
val tyname_to_AST_def = Define‘
29+
Definition tyname_to_AST_def:
3030
tyname_to_AST (Short n) = ND nTyOp [ND nUQTyOp [LF (AlphaT n)]] ∧
3131
tyname_to_AST (Long md (Short n)) = ND nTyOp [LF (LongidT md n)] ∧
3232
tyname_to_AST _ = ARB
33-
’;
33+
End
3434

3535
Theorem tyname_inverted:
3636
∀id. user_expressible_tyname id ⇒

translator/ml_progScript.sml

+4-2
Original file line numberDiff line numberDiff line change
@@ -624,8 +624,10 @@ in
624624
Definition init_env_def[nocompute]:
625625
init_env = ^init_env_tm
626626
End
627-
val init_state_def = Define `
628-
init_state ffi = ^init_state_tm`;
627+
628+
Definition init_state_def:
629+
init_state ffi = ^init_state_tm
630+
End
629631
end
630632

631633
Theorem init_state_env_thm:

0 commit comments

Comments
 (0)