| author | kuncar | 
| Wed, 05 Jun 2013 15:21:52 +0200 | |
| changeset 52307 | 32c433c38ddd | 
| parent 52143 | 36ffe23b25f8 | 
| child 52430 | 289e36c2870a | 
| 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" "|"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
16  | 
and "header" :: diag  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
17  | 
and "chapter" :: thy_heading1  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
18  | 
and "section" :: thy_heading2  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
19  | 
and "subsection" :: thy_heading3  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
20  | 
and "subsubsection" :: thy_heading4  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
21  | 
and "text" "text_raw" :: thy_decl  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
22  | 
and "sect" :: prf_heading2 % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
23  | 
and "subsect" :: prf_heading3 % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
24  | 
and "subsubsect" :: prf_heading4 % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
25  | 
and "txt" "txt_raw" :: prf_decl % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
26  | 
and "classes" "classrel" "default_sort" "typedecl" "type_synonym"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
27  | 
"nonterminal" "arities" "judgment" "consts" "syntax" "no_syntax"  | 
| 51313 | 28  | 
"translations" "no_translations" "defs" "definition"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
29  | 
"abbreviation" "type_notation" "no_type_notation" "notation"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
30  | 
"no_notation" "axiomatization" "theorems" "lemmas" "declare"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
31  | 
"hide_class" "hide_type" "hide_const" "hide_fact" :: thy_decl  | 
| 51295 | 32  | 
and "ML" :: thy_decl % "ML"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
33  | 
and "ML_prf" :: prf_decl % "proof" (* FIXME % "ML" ?? *)  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
34  | 
and "ML_val" "ML_command" :: diag % "ML"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
35  | 
and "setup" "local_setup" "attribute_setup" "method_setup"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
36  | 
"declaration" "syntax_declaration" "simproc_setup"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
37  | 
"parse_ast_translation" "parse_translation" "print_translation"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
38  | 
"typed_print_translation" "print_ast_translation" "oracle" :: thy_decl % "ML"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
39  | 
and "bundle" :: thy_decl  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
40  | 
and "include" "including" :: prf_decl  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
41  | 
and "print_bundles" :: diag  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
42  | 
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
 | 
43  | 
and "sublocale" "interpretation" :: thy_goal  | 
| 
 
c3e99efacb67
back to non-schematic 'sublocale' and 'interpretation' (despite df8fc0567a3d) for more potential parallelism;
 
wenzelm 
parents: 
50603 
diff
changeset
 | 
44  | 
and "interpret" :: prf_goal % "proof"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
45  | 
and "class" :: thy_decl  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
46  | 
and "subclass" :: thy_goal  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
47  | 
and "instantiation" :: thy_decl  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
48  | 
and "instance" :: thy_goal  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
49  | 
and "overloading" :: thy_decl  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
50  | 
and "code_datatype" :: thy_decl  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
51  | 
and "theorem" "lemma" "corollary" :: thy_goal  | 
| 
51274
 
cfc83ad52571
discontinued pointless command category "thy_schematic_goal" -- this is checked dynamically;
 
wenzelm 
parents: 
51270 
diff
changeset
 | 
52  | 
and "schematic_theorem" "schematic_lemma" "schematic_corollary" :: thy_goal  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
53  | 
and "notepad" :: thy_decl  | 
| 
50128
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
49569 
diff
changeset
 | 
54  | 
and "have" :: prf_goal % "proof"  | 
| 
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
49569 
diff
changeset
 | 
55  | 
and "hence" :: prf_goal % "proof" == "then have"  | 
| 
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
49569 
diff
changeset
 | 
56  | 
and "show" :: prf_asm_goal % "proof"  | 
| 
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
49569 
diff
changeset
 | 
57  | 
and "thus" :: prf_asm_goal % "proof" == "then show"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
58  | 
and "then" "from" "with" :: prf_chain % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
59  | 
and "note" "using" "unfolding" :: prf_decl % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
60  | 
and "fix" "assume" "presume" "def" :: prf_asm % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
61  | 
and "obtain" "guess" :: prf_asm_goal % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
62  | 
and "let" "write" :: prf_decl % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
63  | 
and "case" :: prf_asm % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
64  | 
  and "{" :: prf_open % "proof"
 | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
65  | 
and "}" :: prf_close % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
66  | 
and "next" :: prf_block % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
67  | 
and "qed" :: qed_block % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
68  | 
and "by" ".." "." "done" "sorry" :: "qed" % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
69  | 
and "oops" :: qed_global % "proof"  | 
| 
50128
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
49569 
diff
changeset
 | 
70  | 
and "defer" "prefer" "apply" :: prf_script % "proof"  | 
| 
 
599c935aac82
alternative completion for outer syntax keywords;
 
wenzelm 
parents: 
49569 
diff
changeset
 | 
71  | 
and "apply_end" :: prf_script % "proof" == ""  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
72  | 
and "proof" :: prf_block % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
73  | 
and "also" "moreover" :: prf_decl % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
74  | 
and "finally" "ultimately" :: prf_chain % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
75  | 
and "back" :: prf_script % "proof"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
76  | 
and "Isabelle.command" :: control  | 
| 52060 | 77  | 
and "help" "print_commands" "print_options"  | 
| 51585 | 78  | 
"print_context" "print_theory" "print_syntax" "print_abbrevs" "print_defn_rules"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
79  | 
"print_theorems" "print_locales" "print_classes" "print_locale"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
80  | 
"print_interps" "print_dependencies" "print_attributes"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
81  | 
"print_simpset" "print_rules" "print_trans_rules" "print_methods"  | 
| 
49569
 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 
wenzelm 
parents: 
48929 
diff
changeset
 | 
82  | 
"print_antiquotations" "thy_deps" "locale_deps" "class_deps" "thm_deps"  | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
83  | 
"print_binds" "print_facts" "print_cases" "print_statement" "thm"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
84  | 
"prf" "full_prf" "prop" "term" "typ" "print_codesetup" "unused_thms"  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
85  | 
:: diag  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
86  | 
and "cd" :: control  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
87  | 
and "pwd" :: diag  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
88  | 
and "use_thy" "remove_thy" "kill_thy" :: control  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
48638 
diff
changeset
 | 
89  | 
and "display_drafts" "print_drafts" "pr" :: diag  | 
| 
51270
 
17d30843fc3b
reconsider 'pretty_setmargin' as "control" command (instead of "diag") -- it is stateful and Proof General legacy;
 
wenzelm 
parents: 
51224 
diff
changeset
 | 
90  | 
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
 | 
91  | 
and "welcome" :: diag  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48641 
diff
changeset
 | 
92  | 
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
 | 
93  | 
and "end" :: thy_end % "theory"  | 
| 
48646
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48641 
diff
changeset
 | 
94  | 
and "realizers" "realizability" "extract_type" "extract" :: thy_decl  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48641 
diff
changeset
 | 
95  | 
and "find_theorems" "find_consts" :: diag  | 
| 48638 | 96  | 
begin  | 
| 15803 | 97  | 
|
| 48891 | 98  | 
ML_file "Isar/isar_syn.ML"  | 
99  | 
ML_file "Tools/find_theorems.ML"  | 
|
100  | 
ML_file "Tools/find_consts.ML"  | 
|
| 52009 | 101  | 
ML_file "Tools/proof_general_pure.ML"  | 
| 48891 | 102  | 
|
103  | 
||
| 26435 | 104  | 
section {* Further content for the Pure theory *}
 | 
| 20627 | 105  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
106  | 
subsection {* Meta-level connectives in assumptions *}
 | 
| 15803 | 107  | 
|
108  | 
lemma meta_mp:  | 
|
| 18019 | 109  | 
assumes "PROP P ==> PROP Q" and "PROP P"  | 
| 15803 | 110  | 
shows "PROP Q"  | 
| 18019 | 111  | 
by (rule `PROP P ==> PROP Q` [OF `PROP P`])  | 
| 15803 | 112  | 
|
| 23432 | 113  | 
lemmas meta_impE = meta_mp [elim_format]  | 
114  | 
||
| 15803 | 115  | 
lemma meta_spec:  | 
| 26958 | 116  | 
assumes "!!x. PROP P x"  | 
117  | 
shows "PROP P x"  | 
|
118  | 
by (rule `!!x. PROP P x`)  | 
|
| 15803 | 119  | 
|
120  | 
lemmas meta_allE = meta_spec [elim_format]  | 
|
121  | 
||
| 26570 | 122  | 
lemma swap_params:  | 
| 26958 | 123  | 
"(!!x y. PROP P x y) == (!!y x. PROP P x y)" ..  | 
| 26570 | 124  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
125  | 
|
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
126  | 
subsection {* Meta-level conjunction *}
 | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
127  | 
|
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
128  | 
lemma all_conjunction:  | 
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
129  | 
"(!!x. PROP A x &&& PROP B x) == ((!!x. PROP A x) &&& (!!x. PROP B x))"  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
130  | 
proof  | 
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
131  | 
assume conj: "!!x. PROP A x &&& PROP B x"  | 
| 
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
132  | 
show "(!!x. PROP A x) &&& (!!x. PROP B x)"  | 
| 19121 | 133  | 
proof -  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
134  | 
fix x  | 
| 26958 | 135  | 
from conj show "PROP A x" by (rule conjunctionD1)  | 
136  | 
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
 | 
137  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
138  | 
next  | 
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
139  | 
assume conj: "(!!x. PROP A x) &&& (!!x. PROP B x)"  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
140  | 
fix x  | 
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
141  | 
show "PROP A x &&& PROP B x"  | 
| 19121 | 142  | 
proof -  | 
| 26958 | 143  | 
show "PROP A x" by (rule conj [THEN conjunctionD1, rule_format])  | 
144  | 
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
 | 
145  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
146  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
147  | 
|
| 19121 | 148  | 
lemma imp_conjunction:  | 
| 50603 | 149  | 
"(PROP A ==> PROP B &&& PROP C) == ((PROP A ==> PROP B) &&& (PROP A ==> PROP C))"  | 
| 18836 | 150  | 
proof  | 
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
151  | 
assume conj: "PROP A ==> PROP B &&& PROP C"  | 
| 
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
152  | 
show "(PROP A ==> PROP B) &&& (PROP A ==> PROP C)"  | 
| 19121 | 153  | 
proof -  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
154  | 
assume "PROP A"  | 
| 19121 | 155  | 
from conj [OF `PROP A`] show "PROP B" by (rule conjunctionD1)  | 
156  | 
from conj [OF `PROP A`] show "PROP C" by (rule conjunctionD2)  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
157  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
158  | 
next  | 
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
159  | 
assume conj: "(PROP A ==> PROP B) &&& (PROP A ==> PROP C)"  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
160  | 
assume "PROP A"  | 
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
161  | 
show "PROP B &&& PROP C"  | 
| 19121 | 162  | 
proof -  | 
163  | 
from `PROP A` show "PROP B" by (rule conj [THEN conjunctionD1])  | 
|
164  | 
from `PROP A` show "PROP C" by (rule conj [THEN conjunctionD2])  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
165  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
166  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
167  | 
|
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
168  | 
lemma conjunction_imp:  | 
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
169  | 
"(PROP A &&& PROP B ==> PROP C) == (PROP A ==> PROP B ==> PROP C)"  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
170  | 
proof  | 
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
171  | 
assume r: "PROP A &&& PROP B ==> PROP C"  | 
| 22933 | 172  | 
assume ab: "PROP A" "PROP B"  | 
173  | 
show "PROP C"  | 
|
174  | 
proof (rule r)  | 
|
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
175  | 
from ab show "PROP A &&& PROP B" .  | 
| 22933 | 176  | 
qed  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
177  | 
next  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
178  | 
assume r: "PROP A ==> PROP B ==> PROP C"  | 
| 
28856
 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 
wenzelm 
parents: 
28699 
diff
changeset
 | 
179  | 
assume conj: "PROP A &&& PROP B"  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
180  | 
show "PROP C"  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
181  | 
proof (rule r)  | 
| 19121 | 182  | 
from conj show "PROP A" by (rule conjunctionD1)  | 
183  | 
from conj show "PROP B" by (rule conjunctionD2)  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
184  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
185  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
186  | 
|
| 48638 | 187  | 
end  | 
188  |