| author | wenzelm | 
| Mon, 13 Oct 2014 15:45:23 +0200 | |
| changeset 58657 | c917dc025184 | 
| parent 58612 | dbe216a75a4b | 
| child 58660 | 8d4aebb9e327 | 
| permissions | -rw-r--r-- | 
| 
48929
 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 
wenzelm 
parents: 
48891 
diff
changeset
 | 
1  | 
(* Title: Pure/Pure.thy  | 
| 
 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 
wenzelm 
parents: 
48891 
diff
changeset
 | 
2  | 
Author: Makarius  | 
| 
 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 
wenzelm 
parents: 
48891 
diff
changeset
 | 
3  | 
|
| 
 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 
wenzelm 
parents: 
48891 
diff
changeset
 | 
4  | 
Final stage of bootstrapping Pure, based on implicit background theory.  | 
| 
 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 
wenzelm 
parents: 
48891 
diff
changeset
 | 
5  | 
*)  | 
| 
 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 
wenzelm 
parents: 
48891 
diff
changeset
 | 
6  | 
|
| 48638 | 7  | 
theory Pure  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
8  | 
keywords  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
9  | 
    "!!" "!" "%" "(" ")" "+" "," "--" ":" "::" ";" "<" "<=" "=" "=="
 | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
10  | 
"=>" "?" "[" "\<equiv>" "\<leftharpoondown>" "\<rightharpoonup>"  | 
| 52143 | 11  | 
"\<rightleftharpoons>" "\<subseteq>" "]" "and" "assumes"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
12  | 
"attach" "begin" "binder" "constrains" "defines" "fixes" "for"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
13  | 
"identifier" "if" "imports" "in" "includes" "infix" "infixl"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
14  | 
"infixr" "is" "keywords" "notes" "obtains" "open" "output"  | 
| 
51293
 
05b1bbae748d
discontinued obsolete 'uses' within theory header;
 
wenzelm 
parents: 
51274 
diff
changeset
 | 
15  | 
"overloaded" "pervasive" "shows" "structure" "unchecked" "where" "|"  | 
| 
52449
 
79e7fd57acc4
recover command tags from 4cf3f6153eb8 -- important for document preparation;
 
wenzelm 
parents: 
52439 
diff
changeset
 | 
16  | 
and "theory" :: thy_begin % "theory"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
17  | 
and "header" :: diag  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
18  | 
and "chapter" :: thy_heading1  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
19  | 
and "section" :: thy_heading2  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
20  | 
and "subsection" :: thy_heading3  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
21  | 
and "subsubsection" :: thy_heading4  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
22  | 
and "text" "text_raw" :: thy_decl  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
23  | 
and "sect" :: prf_heading2 % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
24  | 
and "subsect" :: prf_heading3 % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
25  | 
and "subsubsect" :: prf_heading4 % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
26  | 
and "txt" "txt_raw" :: prf_decl % "proof"  | 
| 57506 | 27  | 
and "default_sort" :: thy_decl == ""  | 
28  | 
and "typedecl" "type_synonym" "nonterminal" "judgment"  | 
|
| 
55385
 
169e12bbf9a3
discontinued axiomatic 'classes', 'classrel', 'arities';
 
wenzelm 
parents: 
55152 
diff
changeset
 | 
29  | 
"consts" "syntax" "no_syntax" "translations" "no_translations" "defs"  | 
| 
 
169e12bbf9a3
discontinued axiomatic 'classes', 'classrel', 'arities';
 
wenzelm 
parents: 
55152 
diff
changeset
 | 
30  | 
"definition" "abbreviation" "type_notation" "no_type_notation" "notation"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
31  | 
"no_notation" "axiomatization" "theorems" "lemmas" "declare"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
32  | 
"hide_class" "hide_type" "hide_const" "hide_fact" :: thy_decl  | 
| 
56618
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56275 
diff
changeset
 | 
33  | 
and "SML_file" "ML_file" :: thy_load % "ML"  | 
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56275 
diff
changeset
 | 
34  | 
and "SML_import" "SML_export" :: thy_decl % "ML"  | 
| 51295 | 35  | 
and "ML" :: thy_decl % "ML"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
36  | 
and "ML_prf" :: prf_decl % "proof" (* FIXME % "ML" ?? *)  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
37  | 
and "ML_val" "ML_command" :: diag % "ML"  | 
| 
55762
 
27a45aec67a0
suppress completion of obscure keyword, avoid confusion with plain "simp";
 
wenzelm 
parents: 
55516 
diff
changeset
 | 
38  | 
and "simproc_setup" :: thy_decl % "ML" == ""  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
39  | 
and "setup" "local_setup" "attribute_setup" "method_setup"  | 
| 
55762
 
27a45aec67a0
suppress completion of obscure keyword, avoid confusion with plain "simp";
 
wenzelm 
parents: 
55516 
diff
changeset
 | 
40  | 
"declaration" "syntax_declaration"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
41  | 
"parse_ast_translation" "parse_translation" "print_translation"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
42  | 
"typed_print_translation" "print_ast_translation" "oracle" :: thy_decl % "ML"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
43  | 
and "bundle" :: thy_decl  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
44  | 
and "include" "including" :: prf_decl  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
45  | 
and "print_bundles" :: diag  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
46  | 
and "context" "locale" :: thy_decl  | 
| 
51224
 
c3e99efacb67
back to non-schematic 'sublocale' and 'interpretation' (despite df8fc0567a3d) for more potential parallelism;
 
wenzelm 
parents: 
50603 
diff
changeset
 | 
47  | 
and "sublocale" "interpretation" :: thy_goal  | 
| 
 
c3e99efacb67
back to non-schematic 'sublocale' and 'interpretation' (despite df8fc0567a3d) for more potential parallelism;
 
wenzelm 
parents: 
50603 
diff
changeset
 | 
48  | 
and "interpret" :: prf_goal % "proof"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
49  | 
and "class" :: thy_decl  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
50  | 
and "subclass" :: thy_goal  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
51  | 
and "instantiation" :: thy_decl  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
52  | 
and "instance" :: thy_goal  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
53  | 
and "overloading" :: thy_decl  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
54  | 
and "code_datatype" :: thy_decl  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
55  | 
and "theorem" "lemma" "corollary" :: thy_goal  | 
| 
51274
 
cfc83ad52571
discontinued pointless command category "thy_schematic_goal" -- this is checked dynamically;
 
wenzelm 
parents: 
51270 
diff
changeset
 | 
56  | 
and "schematic_theorem" "schematic_lemma" "schematic_corollary" :: thy_goal  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
57  | 
and "notepad" :: thy_decl  | 
| 
50128
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
49569 
diff
changeset
 | 
58  | 
and "have" :: prf_goal % "proof"  | 
| 
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
49569 
diff
changeset
 | 
59  | 
and "hence" :: prf_goal % "proof" == "then have"  | 
| 
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
49569 
diff
changeset
 | 
60  | 
and "show" :: prf_asm_goal % "proof"  | 
| 
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
49569 
diff
changeset
 | 
61  | 
and "thus" :: prf_asm_goal % "proof" == "then show"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
62  | 
and "then" "from" "with" :: prf_chain % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
63  | 
and "note" "using" "unfolding" :: prf_decl % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
64  | 
and "fix" "assume" "presume" "def" :: prf_asm % "proof"  | 
| 
53371
 
47b23c582127
more explicit indication of 'guess' as improper Isar (aka "script") element;
 
wenzelm 
parents: 
52549 
diff
changeset
 | 
65  | 
and "obtain" :: prf_asm_goal % "proof"  | 
| 
 
47b23c582127
more explicit indication of 'guess' as improper Isar (aka "script") element;
 
wenzelm 
parents: 
52549 
diff
changeset
 | 
66  | 
and "guess" :: prf_asm_goal_script % "proof"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
67  | 
and "let" "write" :: prf_decl % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
68  | 
and "case" :: prf_asm % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
69  | 
  and "{" :: prf_open % "proof"
 | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
70  | 
and "}" :: prf_close % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
71  | 
and "next" :: prf_block % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
72  | 
and "qed" :: qed_block % "proof"  | 
| 
53571
 
e58ca0311c0f
more explicit indication of 'done' as proof script element;
 
wenzelm 
parents: 
53371 
diff
changeset
 | 
73  | 
and "by" ".." "." "sorry" :: "qed" % "proof"  | 
| 
 
e58ca0311c0f
more explicit indication of 'done' as proof script element;
 
wenzelm 
parents: 
53371 
diff
changeset
 | 
74  | 
and "done" :: "qed_script" % "proof"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
75  | 
and "oops" :: qed_global % "proof"  | 
| 
50128
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
49569 
diff
changeset
 | 
76  | 
and "defer" "prefer" "apply" :: prf_script % "proof"  | 
| 
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
49569 
diff
changeset
 | 
77  | 
and "apply_end" :: prf_script % "proof" == ""  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
78  | 
and "proof" :: prf_block % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
79  | 
and "also" "moreover" :: prf_decl % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
80  | 
and "finally" "ultimately" :: prf_chain % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
81  | 
and "back" :: prf_script % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
82  | 
and "Isabelle.command" :: control  | 
| 
56069
 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 
wenzelm 
parents: 
55762 
diff
changeset
 | 
83  | 
and "help" "print_commands" "print_options" "print_context"  | 
| 
 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 
wenzelm 
parents: 
55762 
diff
changeset
 | 
84  | 
"print_theory" "print_syntax" "print_abbrevs" "print_defn_rules"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
85  | 
"print_theorems" "print_locales" "print_classes" "print_locale"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
86  | 
"print_interps" "print_dependencies" "print_attributes"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
87  | 
"print_simpset" "print_rules" "print_trans_rules" "print_methods"  | 
| 
56069
 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 
wenzelm 
parents: 
55762 
diff
changeset
 | 
88  | 
"print_antiquotations" "print_ML_antiquotations" "thy_deps"  | 
| 
57415
 
e721124f1b1e
command 'print_term_bindings' supersedes 'print_binds';
 
wenzelm 
parents: 
56864 
diff
changeset
 | 
89  | 
"locale_deps" "class_deps" "thm_deps" "print_binds" "print_term_bindings"  | 
| 
 
e721124f1b1e
command 'print_term_bindings' supersedes 'print_binds';
 
wenzelm 
parents: 
56864 
diff
changeset
 | 
90  | 
"print_facts" "print_cases" "print_statement" "thm" "prf" "full_prf"  | 
| 
 
e721124f1b1e
command 'print_term_bindings' supersedes 'print_binds';
 
wenzelm 
parents: 
56864 
diff
changeset
 | 
91  | 
"prop" "term" "typ" "print_codesetup" "unused_thms" :: diag  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
92  | 
and "use_thy" "remove_thy" "kill_thy" :: control  | 
| 52549 | 93  | 
and "display_drafts" "print_state" "pr" :: diag  | 
| 
52438
 
7b5a5116f3af
back to keyword 'pr' :: diag as required for ProofGeneral command line -- reject this TTY command in Isabelle/jEdit by other means;
 
wenzelm 
parents: 
52437 
diff
changeset
 | 
94  | 
and "pretty_setmargin" "disable_pr" "enable_pr" "commit" "quit" "exit" :: control  | 
| 
48646
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48641 
diff
changeset
 | 
95  | 
and "welcome" :: diag  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48641 
diff
changeset
 | 
96  | 
and "init_toplevel" "linear_undo" "undo" "undos_proof" "cannot_undo" "kill" :: control  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
97  | 
and "end" :: thy_end % "theory"  | 
| 56797 | 98  | 
and "realizers" :: thy_decl == ""  | 
99  | 
and "realizability" :: thy_decl == ""  | 
|
100  | 
and "extract_type" "extract" :: thy_decl  | 
|
| 
48646
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48641 
diff
changeset
 | 
101  | 
and "find_theorems" "find_consts" :: diag  | 
| 
57886
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents: 
57506 
diff
changeset
 | 
102  | 
and "named_theorems" :: thy_decl  | 
| 
52437
 
c88354589b43
more formal ProofGeneral command setup within theory Pure;
 
wenzelm 
parents: 
52430 
diff
changeset
 | 
103  | 
and "ProofGeneral.process_pgip" "ProofGeneral.pr" "ProofGeneral.undo"  | 
| 
 
c88354589b43
more formal ProofGeneral command setup within theory Pure;
 
wenzelm 
parents: 
52430 
diff
changeset
 | 
104  | 
"ProofGeneral.restart" "ProofGeneral.kill_proof" "ProofGeneral.inform_file_processed"  | 
| 
 
c88354589b43
more formal ProofGeneral command setup within theory Pure;
 
wenzelm 
parents: 
52430 
diff
changeset
 | 
105  | 
"ProofGeneral.inform_file_retracted" :: control  | 
| 48638 | 106  | 
begin  | 
| 15803 | 107  | 
|
| 56205 | 108  | 
ML_file "ML/ml_antiquotations.ML"  | 
| 55516 | 109  | 
ML_file "ML/ml_thms.ML"  | 
| 56864 | 110  | 
ML_file "Tools/print_operation.ML"  | 
| 48891 | 111  | 
ML_file "Isar/isar_syn.ML"  | 
| 55141 | 112  | 
ML_file "Isar/calculation.ML"  | 
| 
58544
 
340f130b3d38
bibtex support in ML: document antiquotation @{cite} with markup;
 
wenzelm 
parents: 
58201 
diff
changeset
 | 
113  | 
ML_file "Tools/bibtex.ML"  | 
| 55030 | 114  | 
ML_file "Tools/rail.ML"  | 
| 53707 | 115  | 
ML_file "Tools/rule_insts.ML";  | 
| 
57934
 
5e500c0e7eca
tuned signature -- prefer self-contained user-space tool;
 
wenzelm 
parents: 
57886 
diff
changeset
 | 
116  | 
ML_file "Tools/thm_deps.ML";  | 
| 58201 | 117  | 
ML_file "Tools/class_deps.ML"  | 
| 48891 | 118  | 
ML_file "Tools/find_theorems.ML"  | 
119  | 
ML_file "Tools/find_consts.ML"  | 
|
| 52009 | 120  | 
ML_file "Tools/proof_general_pure.ML"  | 
| 54730 | 121  | 
ML_file "Tools/simplifier_trace.ML"  | 
| 
57886
 
7cae177c9084
support for named collections of theorems in canonical order;
 
wenzelm 
parents: 
57506 
diff
changeset
 | 
122  | 
ML_file "Tools/named_theorems.ML"  | 
| 
58657
 
c917dc025184
support for named plugins for definitional packages;
 
wenzelm 
parents: 
58612 
diff
changeset
 | 
123  | 
ML_file "Tools/plugin.ML"  | 
| 48891 | 124  | 
|
125  | 
||
| 58611 | 126  | 
section \<open>Basic attributes\<close>  | 
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
127  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
128  | 
attribute_setup tagged =  | 
| 58611 | 129  | 
\<open>Scan.lift (Args.name -- Args.name) >> Thm.tag\<close>  | 
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
130  | 
"tagged theorem"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
131  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
132  | 
attribute_setup untagged =  | 
| 58611 | 133  | 
\<open>Scan.lift Args.name >> Thm.untag\<close>  | 
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
134  | 
"untagged theorem"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
135  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
136  | 
attribute_setup kind =  | 
| 58611 | 137  | 
\<open>Scan.lift Args.name >> Thm.kind\<close>  | 
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
138  | 
"theorem kind"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
139  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
140  | 
attribute_setup THEN =  | 
| 58611 | 141  | 
\<open>Scan.lift (Scan.optional (Args.bracks Parse.nat) 1) -- Attrib.thm  | 
142  | 
>> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => A RSN (i, B)))\<close>  | 
|
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
143  | 
"resolution with rule"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
144  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
145  | 
attribute_setup OF =  | 
| 58611 | 146  | 
\<open>Attrib.thms >> (fn Bs => Thm.rule_attribute (fn _ => fn A => A OF Bs))\<close>  | 
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
147  | 
"rule resolved with facts"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
148  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
149  | 
attribute_setup rename_abs =  | 
| 58611 | 150  | 
\<open>Scan.lift (Scan.repeat (Args.maybe Args.name)) >> (fn vs =>  | 
151  | 
Thm.rule_attribute (K (Drule.rename_bvars' vs)))\<close>  | 
|
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
152  | 
"rename bound variables in abstractions"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
153  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
154  | 
attribute_setup unfolded =  | 
| 58611 | 155  | 
\<open>Attrib.thms >> (fn ths =>  | 
156  | 
Thm.rule_attribute (fn context => Local_Defs.unfold (Context.proof_of context) ths))\<close>  | 
|
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
157  | 
"unfolded definitions"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
158  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
159  | 
attribute_setup folded =  | 
| 58611 | 160  | 
\<open>Attrib.thms >> (fn ths =>  | 
161  | 
Thm.rule_attribute (fn context => Local_Defs.fold (Context.proof_of context) ths))\<close>  | 
|
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
162  | 
"folded definitions"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
163  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
164  | 
attribute_setup consumes =  | 
| 58611 | 165  | 
\<open>Scan.lift (Scan.optional Parse.int 1) >> Rule_Cases.consumes\<close>  | 
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
166  | 
"number of consumed facts"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
167  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
168  | 
attribute_setup constraints =  | 
| 58611 | 169  | 
\<open>Scan.lift Parse.nat >> Rule_Cases.constraints\<close>  | 
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
170  | 
"number of equality constraints"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
171  | 
|
| 58611 | 172  | 
attribute_setup case_names =  | 
173  | 
\<open>Scan.lift (Scan.repeat1 (Args.name --  | 
|
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
174  | 
    Scan.optional (@{keyword "["} |-- Scan.repeat1 (Args.maybe Args.name) --| @{keyword "]"}) []))
 | 
| 58611 | 175  | 
>> (fn cs =>  | 
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
176  | 
Rule_Cases.cases_hyp_names  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
177  | 
(map #1 cs)  | 
| 58611 | 178  | 
(map (map (the_default Rule_Cases.case_hypsN) o #2) cs))\<close>  | 
179  | 
"named rule cases"  | 
|
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
180  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
181  | 
attribute_setup case_conclusion =  | 
| 58611 | 182  | 
\<open>Scan.lift (Args.name -- Scan.repeat Args.name) >> Rule_Cases.case_conclusion\<close>  | 
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
183  | 
"named conclusion of rule cases"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
184  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
185  | 
attribute_setup params =  | 
| 58611 | 186  | 
\<open>Scan.lift (Parse.and_list1 (Scan.repeat Args.name)) >> Rule_Cases.params\<close>  | 
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
187  | 
"named rule parameters"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
188  | 
|
| 58611 | 189  | 
attribute_setup rule_format =  | 
190  | 
\<open>Scan.lift (Args.mode "no_asm")  | 
|
191  | 
>> (fn true => Object_Logic.rule_format_no_asm | false => Object_Logic.rule_format)\<close>  | 
|
192  | 
"result put into canonical rule format"  | 
|
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
193  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
194  | 
attribute_setup elim_format =  | 
| 58611 | 195  | 
\<open>Scan.succeed (Thm.rule_attribute (K Tactic.make_elim))\<close>  | 
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
196  | 
"destruct rule turned into elimination rule format"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
197  | 
|
| 58611 | 198  | 
attribute_setup no_vars =  | 
199  | 
\<open>Scan.succeed (Thm.rule_attribute (fn context => fn th =>  | 
|
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
200  | 
let  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
201  | 
val ctxt = Variable.set_body false (Context.proof_of context);  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
202  | 
val ((_, [th']), _) = Variable.import true [th] ctxt;  | 
| 58611 | 203  | 
in th' end))\<close>  | 
204  | 
"imported schematic variables"  | 
|
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
205  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
206  | 
attribute_setup eta_long =  | 
| 58611 | 207  | 
\<open>Scan.succeed (Thm.rule_attribute (fn _ => Conv.fconv_rule Drule.eta_long_conversion))\<close>  | 
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
208  | 
"put theorem into eta long beta normal form"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
209  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
210  | 
attribute_setup atomize =  | 
| 58611 | 211  | 
\<open>Scan.succeed Object_Logic.declare_atomize\<close>  | 
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
212  | 
"declaration of atomize rule"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
213  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
214  | 
attribute_setup rulify =  | 
| 58611 | 215  | 
\<open>Scan.succeed Object_Logic.declare_rulify\<close>  | 
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
216  | 
"declaration of rulify rule"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
217  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
218  | 
attribute_setup rotated =  | 
| 58611 | 219  | 
\<open>Scan.lift (Scan.optional Parse.int 1  | 
220  | 
>> (fn n => Thm.rule_attribute (fn _ => rotate_prems n)))\<close>  | 
|
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
221  | 
"rotated theorem premises"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
222  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
223  | 
attribute_setup defn =  | 
| 58611 | 224  | 
\<open>Attrib.add_del Local_Defs.defn_add Local_Defs.defn_del\<close>  | 
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
225  | 
"declaration of definitional transformations"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
226  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
227  | 
attribute_setup abs_def =  | 
| 58611 | 228  | 
\<open>Scan.succeed (Thm.rule_attribute (fn context =>  | 
229  | 
Local_Defs.meta_rewrite_rule (Context.proof_of context) #> Drule.abs_def))\<close>  | 
|
| 
55140
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
230  | 
"abstract over free variables of definitional theorem"  | 
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
231  | 
|
| 
 
7eb0c04e4c40
define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
 
wenzelm 
parents: 
55030 
diff
changeset
 | 
232  | 
|
| 58611 | 233  | 
section \<open>Further content for the Pure theory\<close>  | 
| 20627 | 234  | 
|
| 58611 | 235  | 
subsection \<open>Meta-level connectives in assumptions\<close>  | 
| 15803 | 236  | 
|
237  | 
lemma meta_mp:  | 
|
| 58612 | 238  | 
assumes "PROP P \<Longrightarrow> PROP Q" and "PROP P"  | 
| 15803 | 239  | 
shows "PROP Q"  | 
| 58612 | 240  | 
by (rule \<open>PROP P \<Longrightarrow> PROP Q\<close> [OF \<open>PROP P\<close>])  | 
| 15803 | 241  | 
|
| 23432 | 242  | 
lemmas meta_impE = meta_mp [elim_format]  | 
243  | 
||
| 15803 | 244  | 
lemma meta_spec:  | 
| 58612 | 245  | 
assumes "\<And>x. PROP P x"  | 
| 26958 | 246  | 
shows "PROP P x"  | 
| 58612 | 247  | 
by (rule \<open>\<And>x. PROP P x\<close>)  | 
| 15803 | 248  | 
|
249  | 
lemmas meta_allE = meta_spec [elim_format]  | 
|
250  | 
||
| 26570 | 251  | 
lemma swap_params:  | 
| 58612 | 252  | 
"(\<And>x y. PROP P x y) \<equiv> (\<And>y x. PROP P x y)" ..  | 
| 26570 | 253  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
254  | 
|
| 58611 | 255  | 
subsection \<open>Meta-level conjunction\<close>  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
256  | 
|
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
257  | 
lemma all_conjunction:  | 
| 58612 | 258  | 
"(\<And>x. PROP A x &&& PROP B x) \<equiv> ((\<And>x. PROP A x) &&& (\<And>x. PROP B x))"  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
259  | 
proof  | 
| 58612 | 260  | 
assume conj: "\<And>x. PROP A x &&& PROP B x"  | 
261  | 
show "(\<And>x. PROP A x) &&& (\<And>x. PROP B x)"  | 
|
| 19121 | 262  | 
proof -  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
263  | 
fix x  | 
| 26958 | 264  | 
from conj show "PROP A x" by (rule conjunctionD1)  | 
265  | 
from conj show "PROP B x" by (rule conjunctionD2)  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
266  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
267  | 
next  | 
| 58612 | 268  | 
assume conj: "(\<And>x. PROP A x) &&& (\<And>x. PROP B x)"  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
269  | 
fix x  | 
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
270  | 
show "PROP A x &&& PROP B x"  | 
| 19121 | 271  | 
proof -  | 
| 26958 | 272  | 
show "PROP A x" by (rule conj [THEN conjunctionD1, rule_format])  | 
273  | 
show "PROP B x" by (rule conj [THEN conjunctionD2, rule_format])  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
274  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
275  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
276  | 
|
| 19121 | 277  | 
lemma imp_conjunction:  | 
| 58612 | 278  | 
"(PROP A \<Longrightarrow> PROP B &&& PROP C) \<equiv> ((PROP A \<Longrightarrow> PROP B) &&& (PROP A \<Longrightarrow> PROP C))"  | 
| 18836 | 279  | 
proof  | 
| 58612 | 280  | 
assume conj: "PROP A \<Longrightarrow> PROP B &&& PROP C"  | 
281  | 
show "(PROP A \<Longrightarrow> PROP B) &&& (PROP A \<Longrightarrow> PROP C)"  | 
|
| 19121 | 282  | 
proof -  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
283  | 
assume "PROP A"  | 
| 58611 | 284  | 
from conj [OF \<open>PROP A\<close>] show "PROP B" by (rule conjunctionD1)  | 
285  | 
from conj [OF \<open>PROP A\<close>] show "PROP C" by (rule conjunctionD2)  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
286  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
287  | 
next  | 
| 58612 | 288  | 
assume conj: "(PROP A \<Longrightarrow> PROP B) &&& (PROP A \<Longrightarrow> PROP C)"  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
289  | 
assume "PROP A"  | 
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
290  | 
show "PROP B &&& PROP C"  | 
| 19121 | 291  | 
proof -  | 
| 58611 | 292  | 
from \<open>PROP A\<close> show "PROP B" by (rule conj [THEN conjunctionD1])  | 
293  | 
from \<open>PROP A\<close> show "PROP C" by (rule conj [THEN conjunctionD2])  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
294  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
295  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
296  | 
|
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
297  | 
lemma conjunction_imp:  | 
| 58612 | 298  | 
"(PROP A &&& PROP B \<Longrightarrow> PROP C) \<equiv> (PROP A \<Longrightarrow> PROP B \<Longrightarrow> PROP C)"  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
299  | 
proof  | 
| 58612 | 300  | 
assume r: "PROP A &&& PROP B \<Longrightarrow> PROP C"  | 
| 22933 | 301  | 
assume ab: "PROP A" "PROP B"  | 
302  | 
show "PROP C"  | 
|
303  | 
proof (rule r)  | 
|
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
304  | 
from ab show "PROP A &&& PROP B" .  | 
| 22933 | 305  | 
qed  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
306  | 
next  | 
| 58612 | 307  | 
assume r: "PROP A \<Longrightarrow> PROP B \<Longrightarrow> PROP C"  | 
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
308  | 
assume conj: "PROP A &&& PROP B"  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
309  | 
show "PROP C"  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
310  | 
proof (rule r)  | 
| 19121 | 311  | 
from conj show "PROP A" by (rule conjunctionD1)  | 
312  | 
from conj show "PROP B" by (rule conjunctionD2)  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
313  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
314  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
315  | 
|
| 48638 | 316  | 
end  | 
317  |